theory law=trajForm + subsetNet +sym : constdefs defNodes ::" entity set => node set" "defNodes enttrs=={n. EX l. (l:enttrs & fanout ( l)=n)}" consts onNodes::" trajForm => node set" primrec "onNodes (Is1 n) = {n}" "onNodes (Is0 n) = {n}" "onNodes (f1 andT f2)= (onNodes f1) Un (onNodes f2)" "onNodes (P ->T tf) =(onNodes tf)" "onNodes ( Next tf) =(onNodes tf)" "onNodes chaos={}" constdefs InducedNetByNames ::" entity set => node set =>entity set" " InducedNetByNames nl nodes0 == subNetList nl {g. EX n. isDefinedIn n nl& n:nodes0 & g=lookUp nl n}" lemma outOfDevIsDefined:"l:nl ==> isDefinedIn (fanout l) nl" apply(unfold isDefinedIn_def) apply auto done lemma nodeNameIsUniqueInNet2: assumes a1:"nl:netlists" and a2:"isDefinedIn n nl" shows " EX! enttr.(enttr):nl & fanout enttr=n" proof - have "EX! enttr.(enttr):nl &(( enttr)= Input n | ( EX data .( enttr)=Delay n data )| (EX inps tab.( enttr)= Gate n inps tab))" (is "EX! enttr. ?P enttr") apply(cut_tac a1 a2) apply(rule nodeNameIsUniqueInNet) by assumption+ then obtain enttr where b1:"?P enttr & (ALL l. ?P l--> l=enttr)" by blast let ?R="λ n. (( enttr)= Input n | ( EX data .( enttr)=Delay n data )| (EX inps tab.( enttr)= Gate n inps tab))" show ?thesis apply(rule_tac a="enttr" in ex1I) apply(cut_tac b1) apply simp apply(erule conjE)+ apply(erule disjE,simp) apply(erule disjE) apply(erule exE,simp) apply((erule exE)+,simp) apply(cut_tac b1,simp) apply(case_tac "x") apply auto done qed lemma isDeifinedInSub:" isDefinedIn m (subNetList nl nl') ==>isDefinedIn m nl " apply (subgoal_tac "subNetList nl nl' ⊆ nl") apply(unfold isDefinedIn_def) apply(auto intro:subsetNetRef) done lemma isDeifinedInSub1:" [|isDefinedIn m ( nl') ; nl'⊆ nl|] ==>isDefinedIn m nl " apply(unfold isDefinedIn_def) apply(auto) done lemma lookUpSub1: assumes a1:"nl' ⊆ nl" and a2:"isDefinedIn n nl'" and a3:"nl:netlists" and a4:"nl': netlists" shows "lookUp nl n=lookUp ( nl') n" proof - let ?R="λ nl enttr. enttr ∈ nl ∧ (enttr = Input n ∨ (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab))" have b1:"EX! l. ?R nl l" apply(rule nodeNameIsUniqueInNet) apply assumption apply(rule isDeifinedInSub1) by assumption have b2:"EX! l. ?R ( nl') l" apply(rule nodeNameIsUniqueInNet) apply(cut_tac a4) apply assumption+ done then obtain l where b3 :"?R ( nl') l &(∀ l0. ?R ( nl') l0 --> l0=l)" by blast then obtain b4:"?R nl l " apply(cut_tac a1 a4,blast) done then have b5:" (∀ l0. ?R nl l0 --> l0=l)" apply(cut_tac b1) apply blast done from b3 have b6:"lookUp nl' n=l" apply - apply(rule lookUpIsTheDev ) apply(cut_tac a4) apply assumption+ apply auto done from b4 b5 have b7:"lookUp nl n=l" apply - apply(rule lookUpIsTheDev ) apply assumption+ done with b6 show ?thesis by simp qed lemma lookUpSub: assumes a1:"nl' ⊆ nl" and a2:"isDefinedIn n (subNetList nl nl')" and a3:"nl:netlists" shows "lookUp nl n=lookUp (subNetList nl nl') n" proof - let ?R="λ nl enttr. enttr ∈ nl ∧ (enttr = Input n ∨ (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab))" have b1:"EX! l. ?R nl l" apply(rule nodeNameIsUniqueInNet) apply assumption apply(rule isDeifinedInSub) by assumption have b2:"EX! l. ?R (subNetList nl nl') l" apply(rule nodeNameIsUniqueInNet) apply(rule subNetIsNetlist) apply assumption+ done then obtain l where b3 :"?R (subNetList nl nl') l &(∀ l0. ?R (subNetList nl nl') l0 --> l0=l)" by blast then obtain b4:"?R nl l " apply(subgoal_tac "(subNetList nl nl') ⊆ nl") apply(cut_tac a1,blast) apply(rule subsetI) apply(rule subsetNetRef) apply assumption+ done then have b5:" (∀ l0. ?R nl l0 --> l0=l)" apply(cut_tac b1) apply blast done from b3 have b6:"lookUp (subNetList nl nl') n=l" apply - apply(rule lookUpIsTheDev ) apply(rule subNetIsNetlist) apply assumption+ apply auto done from b4 b5 have b7:"lookUp nl n=l" apply - apply(rule lookUpIsTheDev ) apply assumption+ done with b6 show ?thesis by simp qed lemma lookUpDevName:"[|nl:netlists; enttr:nl|]==> (lookUp ( nl ) (fanout enttr)) =enttr " apply(rule lookUpIsTheDev) apply simp apply(case_tac "enttr") apply auto done lemma lookUpFanOut: assumes a1:"nl:netlists" and a2:"isDefinedIn n nl" shows "fanout (lookUp ( nl ) n) =n" proof - have b1:"EX! enttr.(enttr):nl & fanout enttr=n"(is "EX! l. ?P l") proof(rule nodeNameIsUniqueInNet2)qed then obtain a where b2:"?P (a::entity)&(ALL l. ?P l --> l=a)" by blast show "fanout (lookUp ( nl ) n) =n" thm theI2 proof(unfold lookUp_def,rule_tac a="(a::entity)" in theI2) show "(a )∈ nl ∧ (a = Input n ∨ (∃data. a = Delay n data) ∨ (∃inps tab. a = Gate n inps tab))" proof(cut_tac b2,case_tac "a::entity",force+)qed next fix x assume c1:" x ∈ nl ∧ (x = Input n ∨ (∃data. x = Delay n data) ∨ (∃inps tab. x = Gate n inps tab))" show "x=a" proof(cut_tac c1 b2,simp,force)qed qed(auto) qed lemma subsetNetIsClosed: assumes a1:"nl: netlists" and a2:"isClosed nl" and a3:"nl' ⊆ nl" shows "isClosed (subNetList nl nl')" proof(unfold isClosed_def,(rule allI)+,(rule impI)+) fix m n assume b1:"isDefinedIn m (subNetList nl nl')" and b2:" n ∈ set (fanins (lookUp (subNetList nl nl') m))" from b1 have b1a:"isDefinedIn m nl" by(rule isDeifinedInSub) have "lookUp (subNetList nl nl') m = (lookUp nl m)" apply (rule sym) apply( rule lookUpSub) by assumption+ from this b2 have b2a:" n ∈ set (fanins (lookUp nl m))" by simp from b1 have b3:"EX! l. l∈subNetList nl nl' & fanout l = m " (is "EX! l. ?P l") thm nodeNameIsUniqueInNet2 apply - apply(rule nodeNameIsUniqueInNet2) apply(rule subNetIsNetlist) by assumption+ (*proof(unfold isDefinedIn_def, blast) qed*) then obtain l where b4:"?P l" by auto then have "m=fanout l" by simp then have "(lookUp nl m)= l" apply - apply simp apply(rule lookUpDevName) apply assumption apply(cut_tac b4) apply(erule conjE) apply(rule subsetNetRef) by assumption with b2a have b5:"n ∈ set (fanins l)" by simp have "isDefinedIn n nl" apply(cut_tac a2) apply(unfold isClosed_def) apply(drule_tac x="m" in spec, drule_tac x="n" in spec) apply(cut_tac b1a b2a) apply simp done then have "EX enttr1. enttr1:nl & (fanout enttr1)=n " (is "EX enttr1. ?Q enttr1") apply - apply(unfold isDefinedIn_def) by auto then obtain enttr1 where b8:"?Q enttr1" by auto then have "enttr1:subNetList nl nl'" apply(cut_tac b4) apply(cut_tac b5) apply(rule_tac ?enttr0.0="l" and ?enttr1.0="enttr1" in subAddLink) apply auto done then show "isDefinedIn n (subNetList nl nl')" apply - apply(unfold isDefinedIn_def) apply(cut_tac b8) apply auto done qed lemma independentAux1: shows " ALL sq sq0. (ALL t n. n∈ onNodes A --> sq t n = sq0 t n) --> (sq \<Turnstile> A) = (sq0 \<Turnstile>A)" (is "ALL sq sq0. ?ant sq sq0 A --> ?cons sq sq0 A") proof(induct_tac A) fix n show "ALL sq sq0. ?ant sq sq0 (Is1 n) --> ?cons sq sq0 (Is1 n)" apply auto done next fix n show "ALL sq sq0. ?ant sq sq0 (Is0 n) --> ?cons sq sq0 (Is0 n)" apply auto done next fix trajForm1 trajForm2 assume b1:"ALL sq sq0.?ant sq sq0 (trajForm1) -->?cons sq sq0 (trajForm1)" and b2:"ALL sq sq0.?ant sq sq0 (trajForm2) -->?cons sq sq0 (trajForm2)" let ?f="(trajForm1 andT trajForm2)" show "ALL sq sq0.?ant sq sq0 (?f) -->?cons sq sq0 (?f)" apply(cut_tac b1 b2) apply(rule allI)+ apply(rule impI) apply(drule_tac x="sq" in spec) apply(drule_tac x="sq" in spec) apply(drule_tac x="sq0" in spec)+ apply auto done next fix P trajForm2 assume b1:"ALL sq sq0.?ant sq sq0 (trajForm2) -->?cons sq sq0 (trajForm2)" let ?f="(P->TtrajForm2)" show "ALL sq sq0. ?ant sq sq0 (?f) -->?cons sq sq0 (?f)" apply(cut_tac b1 , auto) done next fix trajForm assume b1:"ALL sq sq0. ?ant sq sq0 (trajForm) -->?cons sq sq0 (trajForm)" let ?f="( Next trajForm)" show "ALL sq sq0. ?ant sq sq0 (?f) -->?cons sq sq0 (?f)" apply(cut_tac b1 , simp) apply(unfold suffix_def,simp) done next let ?f="chaos" show "ALL sq sq0. ?ant sq sq0 (?f) -->?cons sq sq0 (?f)" by simp qed lemma independentAux2: shows "ALL t n. n ~:(onNodes A) --> (defSqOfTrForm A) t n=X" (is "ALL t n. ?ant n A --> ?cons A t n") proof(induct_tac A) fix m let ?A="Is1 m" show "ALL t n. ?ant n ?A --> ?cons ?A t n" proof((rule allI)+,rule impI) fix t n assume b1:"?ant n ?A" show "?cons ?A t n" by(cut_tac b1,auto) qed next fix m let ?A="Is0 m" show "ALL t n. ?ant n ?A --> ?cons ?A t n" proof((rule allI)+,rule impI) fix t n assume b1:"?ant n ?A" show "?cons ?A t n" by(cut_tac b1,auto) qed next fix trajForm1 trajForm2 assume b1:"ALL t n. ?ant n trajForm1--> ?cons trajForm1 t n" and b2:"ALL t n. ?ant n trajForm2--> ?cons trajForm2 t n" let ?A="(trajForm1 andT trajForm2)" show "ALL t n. ?ant n ?A --> ?cons ?A t n" proof((rule allI)+,rule impI) fix t n assume c1:"?ant n ?A" show "?cons ?A t n" proof(cut_tac b1 b2 c1,auto intro:lubId) qed qed next fix P trajForm1 assume b1:"ALL t n. ?ant n trajForm1--> ?cons trajForm1 t n" let ?A="(P ->T trajForm1)" show "ALL t n. ?ant n ?A --> ?cons ?A t n" proof((rule allI)+,rule impI) fix t n assume c1:"?ant n ?A" show "?cons ?A t n" proof(cut_tac b1 c1,case_tac "P=True",auto,unfold X_def,simp) qed qed next fix trajForm1 assume b1:"ALL t n. ?ant n trajForm1--> ?cons trajForm1 t n" let ?A="(Next trajForm1)" show "ALL t n. ?ant n ?A --> ?cons ?A t n" proof((rule allI)+,rule impI) fix t n assume c1:"?ant n ?A" show "?cons ?A t n" proof(cut_tac b1 c1,auto) qed qed next let ?A="chaos" show "ALL t n. ?ant n ?A --> ?cons ?A t n" by simp qed lemma independentAux2: shows " n ~:(onNodes A) ==> (defSqOfTrForm A) t n=X" apply(cut_tac independentAux2[where A="A"]) apply blast done thm fSeqIsMonoAny lemma independentAux3: assumes a1:"nl:netlists" and a2:"ALL t n. isDefinedIn n nl --> sq0 t n =sq t n" and a3:"isClosed nl" shows " ALL t n. isDefinedIn n nl --> (fSeq nl sq0) t n= (fSeq nl sq ) t n" (is "?P t") proof((rule allI)+, rule impI) fix t n assume b1:"isDefinedIn n nl " have b2:"(fSeq nl sq0) t n \<sqsubseteq> (fSeq nl sq ) t n" proof(rule fSeqIsMonoAny) show "∀t n. isDefinedIn n nl --> sq0 t n \<sqsubseteq> sq t n" proof(cut_tac b1 a2, auto intro:leqReflexive) qed qed have b3:"(fSeq nl sq) t n \<sqsubseteq> (fSeq nl sq0 ) t n" proof(rule fSeqIsMonoAny) show "∀t n. isDefinedIn n nl --> sq t n \<sqsubseteq> sq0 t n" proof(cut_tac b1 a2, auto intro:leqReflexive) qed qed with b2 show "fSeq nl sq0 t n = fSeq nl sq t n" proof(rule eqIByLeq)qed qed lemma XIsLubUnit: "lub a X=a" apply(unfold lub_def) apply(unfold X_def) by auto lemma rclosureSubNet: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:" x: (rclosure nl0 s ) " shows " x:(rclosure nl s ) " apply(cut_tac a1 a2 a3 a4) apply(erule rclosure.induct) apply(auto intro:stAddInput stAddDelay stAddGate) done lemma pairInzipList: "ALL stateLs. pair mem zip inps stateLs -->(fst pair) mem inps" apply(induct inps ) apply simp apply(rule allI,rule impI) apply(case_tac stateLs) apply simp apply(subgoal_tac "pair=(a,aa)|pair mem zip inps list") apply(erule disjE) apply simp apply(drule_tac x="list" in spec,simp) apply(thin_tac " ∀stateLs. pair mem zip inps stateLs --> fst pair mem inps") apply simp apply(case_tac "(a,aa)=pair") apply simp apply simp done lemma pairInzipList: "pair mem zip inps stateLs ==>(fst pair) mem inps" apply(cut_tac pairInzipList[where pair="pair" and inps="inps" ]) apply blast done lemma GateInNetlist: assumes a2:"nl:netlists" and a3:" m mem inps" shows "Gate n inps tab :nl --> isDefinedIn m nl" apply - apply(cut_tac a2 a3) apply(erule netlists.induct) apply(unfold isDefinedIn_def, auto) done lemma GateInNetlist: assumes a1:"Gate n inps tab :nl" and a2:"nl:netlists" and a3:" m mem inps" shows " isDefinedIn m nl" apply(cut_tac GateInNetlist[where nl="nl" and m="m" and inps="inps" and n="n" and tab="tab"]) by(cut_tac a1,auto dest: GateInNetlist) lemma rclosureSubNet': assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" shows "( x: (rclosure nl s ) ) --> isDefinedIn (fst x) nl0 --> (x:(rclosure nl0 s )) " proof assume b1:" x ∈ rclosure nl s" show "isDefinedIn (fst x) nl0 -->x ∈ rclosure nl0 s" (is "?ant x -->?cons x") using b1 proof(induct) fix xa let ?dev="Input xa" let ?x="(xa, s xa)" assume c1:"?dev ∈ nl" show "?ant ?x -->?cons ?x" proof assume d1:"?ant ?x " have "EX l. l:nl0 & fanout l= xa" (is "EX l. ?R l nl0") apply(cut_tac c1 d1) apply(unfold isDefinedIn_def) apply simp apply blast done then obtain l where d2:"?R l nl0" by blast then have b3:"?R l nl" by(cut_tac a3,blast) from c1 and this have "l=?dev" apply(cut_tac a1 ) apply(drule_tac n="xa" in nodeNameIsUniqueInNet2) apply(cut_tac d1 a3) apply(unfold isDefinedIn_def,auto) done with d2 show " ?x :rclosure nl0 s" apply - apply simp apply(auto intro:stAddInput) done qed next fix data n let ?dev="Delay n data" let ?x="(n, s n)" assume c1:"?dev ∈ nl" show "?ant ?x -->?cons ?x" proof assume d1:"?ant ?x " have "EX l. l:nl0 & fanout l= n" (is "EX l. ?R l nl0") apply(cut_tac c1 d1) apply(unfold isDefinedIn_def) apply simp apply blast done then obtain l where d2:"?R l nl0" by blast then have b3:"?R l nl" by(cut_tac a3,blast) from c1 and this have "l=?dev" apply(cut_tac a1 ) apply(drule_tac n="n" in nodeNameIsUniqueInNet2) apply(cut_tac d1 a3) apply(unfold isDefinedIn_def,auto) done with d2 show " ?x :rclosure nl0 s" apply - apply simp apply(auto intro:stAddDelay) done qed next fix inps n stateLs tab let ?dev="Gate n inps tab" let ?x="(n,lub (funOfTab tab stateLs) (s n))" assume c1:"?dev ∈ nl" and c2:"length stateLs = length inps" and c3:"∀l. l ∈ set tab --> length l = length inps" and c4:" ∀pair0. pair0 mem zip inps stateLs --> pair0 ∈ rclosure nl s ∧ (isDefinedIn (fst pair0) nl0 --> pair0 ∈ rclosure nl0 s)" show "?ant ?x -->?cons ?x" proof assume d1:"?ant ?x " have "EX l. l:nl0 & fanout l= n" (is "EX l. ?R l nl0") apply(cut_tac c1 d1) apply(unfold isDefinedIn_def) apply simp apply blast done then obtain l where d2:"?R l nl0" by blast then have b3:"?R l nl" by(cut_tac a3,blast) from c1 and this have d2a:"l=?dev" apply(cut_tac a1 ) apply(drule_tac n="n" in nodeNameIsUniqueInNet2) apply(cut_tac d1 a3) apply(unfold isDefinedIn_def,auto) done with d2 show " ?x :rclosure nl0 s" proof(rule_tac stAddGate,simp) show "length stateLs = length inps" by assumption show "∀l. l ∈ set tab --> length l = length inps" by assumption show "∀pair0. pair0 mem zip inps stateLs --> pair0 ∈ rclosure nl0 s" proof(rule allI,rule impI) fix pair0 assume e1:"pair0 mem zip inps stateLs " obtain x y where e2:"pair0=(x,y)"thm PairE apply( erule PairE) done from e1 have "(fst pair0) mem inps" by(rule pairInzipList) then have "isDefinedIn (fst pair0) nl0" apply - apply(rule GateInNetlist) apply(cut_tac d2 d2a,simp) by assumption+ with e1 c4 show "pair0 ∈ rclosure nl0 s" by blast qed qed qed qed qed lemma rclosureSubNet': assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:" x: (rclosure nl s ) " and a5:"isDefinedIn (fst x) nl0 " shows " (x:(rclosure nl0 s )) " apply (cut_tac rclosureSubNet'[where nl="nl" and ?nl0.0="nl0" and s="s" and x="x"] ) apply(cut_tac a4 a5,simp) apply assumption+ done lemma fclosureSubNet: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isDefinedIn n nl0" shows " fclosure nl s n =fclosure nl0 s n" proof - have b1:"(n,fclosure nl s n):rclosure nl s" apply(rule fclosureIsrclosure) apply assumption apply(cut_tac a4 a3) apply(unfold isDefinedIn_def) apply auto done have b2:"(n,fclosure nl0 s n):rclosure nl0 s" apply(rule fclosureIsrclosure) apply assumption apply(cut_tac a4) by assumption have b3:"(n,fclosure nl s n):rclosure nl0 s" apply(rule_tac rclosureSubNet') apply assumption+ apply(cut_tac b1,simp) apply(cut_tac a4,simp) done have b4:"(n,fclosure nl0 s n):rclosure nl s" thm rclosureSubNet apply(rule_tac ?nl0.0="nl0" in rclosureSubNet) apply assumption+ apply(cut_tac b2,simp) done from b3 show ?thesis apply - apply(rule sym) apply(cut_tac a4 a2 ) apply(subgoal_tac "fst (n,fclosure nl s n)=n") apply(drule_tac nl="nl0" and n="n" and pair="(n,fclosure nl s n)" in fclosureVal) apply auto done qed lemma curFunEqAny: assumes a1:"nl:netlists" and a2:" ALL n. ( isDefinedIn n nl)--> (s1 n)= (s2 n)" and a3:"( isDefinedIn n nl)" shows " fclosure nl s1 n =fclosure nl s2 n" proof - have b2:"fclosure nl s1 n \<sqsubseteq> fclosure nl s2 n" proof(rule curFunMonoAny) show "∀ n. isDefinedIn n nl --> s1 n \<sqsubseteq> s2 n" proof(cut_tac a2, auto intro:leqReflexive) qed qed have b3:"fclosure nl s2 n \<sqsubseteq> fclosure nl s1 n" proof(rule curFunMonoAny) show "∀ n. isDefinedIn n nl --> s2 n \<sqsubseteq> s1 n" proof(cut_tac a2, auto intro:leqReflexive) qed qed with b2 show "fclosure nl s1 n = fclosure nl s2 n" proof(rule eqIByLeq)qed qed lemma fSeqSubNet: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isClosed nl0" shows "ALL n. isDefinedIn n nl0 --> fSeq nl sq t n =fSeq nl0 sq t n"(is "?P t") proof (induct_tac t) show "?P 0" apply simp apply(rule allI) apply(rule impI) apply(rule fclosureSubNet) by assumption+ next fix t assume b1:"?P t" show "?P (Suc t)" proof(rule allI,rule impI,simp add:Let_def ) fix n assume b2:"isDefinedIn n nl0" let ?s="(λx. if isDelayNames nl x then let l = lookUp nl x; inps = fanins l; v1 = lub (sq (Suc t) x) (fSeq nl sq t (hd inps)) in v1 else sq (Suc t) x)" let ?s'="(λx. if isDelayNames nl0 x then let l = lookUp nl0 x; inps = fanins l; v1 = lub (sq (Suc t) x) (fSeq nl0 sq t (hd inps)) in v1 else sq (Suc t) x)" have b3:"ALL n. isDefinedIn n nl0 --> ?s n=?s' n" proof(rule allI,rule impI) fix m assume c1:"isDefinedIn m nl0" have "isDelayNames nl0 m |~isDelayNames nl0 m" by simp moreover { assume c2:"isDelayNames nl0 m" from c2 have c3:"isDelayNames nl m" apply - apply(cut_tac a3) apply(unfold isDelayNames_def) apply auto done from c2 obtain l data where b3a:"Delay m data = l ∧ l ∈ nl0" apply(unfold isDelayNames_def) apply auto done then have b4a: "lookUp nl0 m=Delay m data " apply - apply(rule lookUpIsTheDev) apply assumption apply auto done then have c4:"isDefinedIn (hd (fanins (lookUp nl0 m))) nl0" apply - apply simp apply(cut_tac c2) apply(cut_tac a4) apply(unfold isClosed_def) apply(cut_tac c1 ) apply(drule_tac x="m" in spec) apply(drule_tac x="data" in spec) apply simp done have c4a:"lookUp nl m= lookUp nl0 m" by(rule lookUpSub1) have c5: "(hd (fanins (lookUp nl m)))=(hd (fanins (lookUp nl0 m)))" apply(cut_tac c4a b4a) by auto from c2 c3 have "?s m= ?s' m" apply(simp add:Let_def) apply(cut_tac c5,simp) apply(subgoal_tac "fSeq nl sq t (hd (fanins (lookUp nl0 m)))= fSeq nl0 sq t (hd (fanins (lookUp nl0 m)))") apply simp apply(cut_tac c4 b1) by blast } moreover { assume c2:"~isDelayNames nl0 m" from a1 a3 c2 have c3:"~isDelayNames nl m" apply - apply(rule ccontr) apply(cut_tac c1) apply(unfold isDefinedIn_def) apply(erule bexE) apply(case_tac l) apply simp apply(unfold isDelayNames_def) apply auto apply(auto intro:InputNameDelayNameIsDiff DelayNameGateNameIsDiff) done with c2 have "?s m= ?s' m" by auto } ultimately show "?s m= ?s' m" by blast qed then have " fclosure nl ?s n= fclosure nl0 ?s n" apply - apply(rule fclosureSubNet) by assumption+ also have "…=fclosure nl0 ?s' n" apply(rule curFunEqAny) apply assumption apply(cut_tac b3,simp) by assumption finally show "fclosure nl ?s n=fclosure nl0 ?s' n" by auto qed qed lemma fSeqSubNet: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isClosed nl0" and a5:"isDefinedIn n nl0" shows " fSeq nl sq t n =fSeq nl0 sq t n"(is "?P t") apply(cut_tac fSeqSubNet[where nl="nl" and ?nl0.0="nl0" ]) by(cut_tac a1 a2 a3 a4 a5,blast ) lemma trajOfFormInSubNet: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isClosed nl0" shows "ALL n. isDefinedIn n nl0 --> defTrajOfCirc (B ) nl t n =defTrajOfCirc B nl0 t n"(is "ALL n. ?P t n") proof(unfold defTrajOfCirc_def,rule allI,rule impI) fix n assume b1:" isDefinedIn n nl0" show "fSeq nl (defSqOfTrForm B) t n = fSeq nl0 (defSqOfTrForm B) t n" proof(rule fSeqSubNet)qed qed lemma trajOfFormInSubNet: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isClosed nl0" and a5:"isDefinedIn n nl0" shows " defTrajOfCirc (B ) nl t n =defTrajOfCirc B nl0 t n" apply(cut_tac trajOfFormInSubNet[where nl="nl" and ?nl0.0="nl0" ]) by(cut_tac a1 a2 a3 a4 a5,blast ) lemma steSubsetI: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isClosed nl0" and a6:"isClosed nl" and a5:"ALL n. n:( onNodes C) --> isDefinedIn n nl0" shows "(cktSat nl (A \<leadsto> C))=(cktSat nl0 (A \<leadsto> C))" proof assume b1:"cktSat nl (A \<leadsto> C)" show "cktSat nl0 (A \<leadsto> C)" proof(rule fundmental1) have b2:"defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc A nl" proof(cut_tac a4,rule fundmental2)qed show "defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc A nl0" proof(unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n show " defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl0 t n" proof - have "n:(onNodes C) | n~:(onNodes C)" by simp moreover {assume c1:"n:(onNodes C)" from c1 a5 have c2:" isDefinedIn n nl0" by blast from b2 c2 have c3:"defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl t n" by (unfold stateSqLeq_def stateLeq_def ,auto) have c4:" defTrajOfCirc (A ) nl t n =defTrajOfCirc A nl0 t n" proof(cut_tac c2,rule trajOfFormInSubNet)qed with c3 have "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl0 t n" by simp } moreover {assume c1:"~n:(onNodes C)" have "defSqOfTrForm C t n=X" proof(cut_tac c1,rule independentAux2)qed then have "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl0 t n" thm XIsLeastValue proof(simp,rule_tac XIsLeastValue)qed } ultimately show "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl0 t n" by blast qed qed qed next assume b1:"cktSat nl0 (A \<leadsto> C)" show "cktSat nl (A \<leadsto> C)" proof(rule fundmental1) have b2:"defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc A nl0" proof(cut_tac a4,rule fundmental2)qed show "defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc A nl" proof(unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n show " defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl t n" proof - have "n:(onNodes C) | n~:(onNodes C)" by simp moreover {assume c1:"n:(onNodes C)" from c1 a5 have c2:" isDefinedIn n nl0" by blast from b2 c2 have c3:"defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl0 t n" by (unfold stateSqLeq_def stateLeq_def ,auto) have c4:" defTrajOfCirc (A ) nl0 t n =defTrajOfCirc A nl t n" proof(cut_tac c2,rule sym,rule trajOfFormInSubNet)qed with c3 have "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl t n" by simp } moreover {assume c1:"~n:(onNodes C)" have "defSqOfTrForm C t n=X" proof(cut_tac c1,rule independentAux2)qed then have "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl t n" thm XIsLeastValue proof(simp,rule_tac XIsLeastValue)qed } ultimately show "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl t n" by blast qed qed qed qed lemma subsetSubNet2: assumes a1:"nl:netlists" and a2:"nl0 ⊆ nl" shows "nl0 ⊆ subNetList nl nl0" apply(rule subsetI) apply(rule subAddself) apply assumption apply(cut_tac a2,blast) done lemma independentNodes: assumes a1:"(onNodes A) Int ( defNodes (InducedNetByNames nl (onNodes C)))={}" and a2:"isClosed nl" and a3:"nl:netlists" and a4:"ALL n. n:(onNodes C) --> isDefinedIn n nl" shows "(cktSat nl ((B andT A) \<leadsto> C)) = (cktSat nl (B \<leadsto> C)) " proof assume b1:"cktSat nl ((B andT A )\<leadsto> C)" thm fundmental2 show " (cktSat nl (B \<leadsto> C))" proof(rule_tac fundmental1,assumption) prefer 2 show "isClosed nl" by assumption next show " defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc B nl" proof(unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n let ?nl="InducedNetByNames nl (onNodes C)" let ?C="nodes ?nl" have c0a:"{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n} ⊆ nl" apply(rule subsetI) apply simp apply(erule exE) apply(unfold isDefinedIn_def) thm conjE apply(erule conjE)+ apply(erule bexE) apply(cut_tac a3) apply(drule lookUpDevName) apply simp apply auto done have c1a: "InducedNetByNames nl (onNodes C) ∈ netlists" proof(unfold InducedNetByNames_def, rule subNetIsNetlist, assumption, cut_tac c0a, simp) qed have d1a:" InducedNetByNames nl (onNodes C) ⊆ nl" apply(unfold InducedNetByNames_def) apply(blast intro:subsetNetRef) done have d1b:" isClosed (InducedNetByNames nl (onNodes C))" apply(unfold InducedNetByNames_def) apply(rule subsetNetIsClosed) apply assumption+ apply(cut_tac c0a,simp) done show " defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc B nl t n" proof(case_tac " isDefinedIn n ?nl") assume c1:"isDefinedIn n ?nl" have b3:"ALL t n. isDefinedIn n ?nl --> (defTrajOfCirc B ?nl t n )= ( defTrajOfCirc (B andT A ) ?nl t n)" proof(unfold defTrajOfCirc_def,rule independentAux3 ) from c1a show b3a:"InducedNetByNames nl (onNodes C) ∈ netlists" by simp next show "∀t n. isDefinedIn n (InducedNetByNames nl (onNodes C)) --> defSqOfTrForm B t n = defSqOfTrForm (B andT A) t n" proof((rule allI)+,rule impI) fix ta na assume d1:"isDefinedIn na (InducedNetByNames nl (onNodes C))" show "defSqOfTrForm B ta na = defSqOfTrForm (B andT A) ta na" proof - have d2:" ~na :onNodes A" apply(cut_tac a1 d1) apply(unfold defNodes_def) apply(unfold isDefinedIn_def) apply blast done then have d3:"(defSqOfTrForm (A ) ta na ) =X" apply(cut_tac d1) thm independentAux2 apply(rule independentAux2) by assumption have "defSqOfTrForm (B andT A) ta na= lub (defSqOfTrForm (B ) ta na ) (defSqOfTrForm (A ) ta na )" by simp also have "…= (defSqOfTrForm (B ) ta na )" apply(cut_tac d3,simp add:XIsLubUnit) done finally have "defSqOfTrForm (B andT A) ta na=(defSqOfTrForm (B ) ta na )" by simp then show "defSqOfTrForm B ta na = defSqOfTrForm (B andT A) ta na" by simp qed qed next show " isClosed (InducedNetByNames nl (onNodes C))" proof(unfold InducedNetByNames_def,rule subsetNetIsClosed) show "{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n} ⊆ nl" proof(unfold isDefinedIn_def,auto) fix l assume b1:"l:nl" show "lookUp nl (fanout l) ∈ nl" proof - have c1: "lookUp nl (fanout l)=l" proof(cut_tac b1, rule lookUpDevName) qed with b1 show "lookUp nl (fanout l) ∈ nl" by simp qed qed qed qed have b2:"(defSqOfTrForm C) \<sqsubseteq>sq ( defTrajOfCirc (B andT A ) nl )" proof(rule_tac fundmental2,assumption+) qed then have "(defSqOfTrForm C) t n \<sqsubseteq> defTrajOfCirc(B andT A ) nl t n " apply(cut_tac b2, unfold stateSqLeq_def stateLeq_def, auto) done also have " …= ( defTrajOfCirc (B andT A ) ?nl t n)" apply(unfold defTrajOfCirc_def) apply(rule fSeqSubNet) apply assumption apply(cut_tac c1a,simp) apply(cut_tac d1a,simp) apply(cut_tac d1b,simp) by assumption also have "…= ( defTrajOfCirc (B ) ?nl t n)" apply(cut_tac c1 b3, auto) done also have "…=( defTrajOfCirc (B ) nl t n)" apply(rule sym) apply(unfold defTrajOfCirc_def) apply(rule fSeqSubNet) apply assumption apply(cut_tac c1a,simp) apply(cut_tac d1a,simp) apply(cut_tac d1b,simp) by assumption finally show "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc B nl t n" by auto next assume c1:"¬ isDefinedIn n ?nl" let ?S=" {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n}" have "(n: onNodes C) --> isDefinedIn n ?S" apply(rule impI) apply(cut_tac a4) apply(unfold isDefinedIn_def) apply(drule_tac x="n" in spec) apply simp apply(erule bexE) apply(rule_tac x="l" in exI) apply simp apply(rule_tac x="n" in exI) apply( auto intro:lookUpIsTheDev lookUpDevName) apply(rule sym) apply(rule lookUpIsTheDev) apply assumption apply(case_tac l) apply auto done then have "(n: onNodes C) --> isDefinedIn n ?nl" apply - apply(rule impI) apply simp apply(rule isDeifinedInSub1) apply simp apply(unfold InducedNetByNames_def) apply(rule subsetSubNet2) apply assumption apply(cut_tac c0a) apply simp done with c1 have "¬ n ∈ onNodes C" apply - apply(rule ccontr) by blast then have "defSqOfTrForm C t n=X" apply(rule independentAux2) done then show " defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc B nl t n" by( simp add:XIsLeastValue) qed qed qed next assume b1:"cktSat nl (B \<leadsto> C) " show " cktSat nl (B andT A \<leadsto> C)" proof(rule_tac fundmental1,assumption) prefer 2 show "isClosed nl" by assumption next show "defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc (B andT A) nl" proof - have b2:"defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc B nl" proof( cut_tac b1,rule_tac fundmental2) qed also have "…\<sqsubseteq>sq defTrajOfCirc(B andT A) nl" proof(unfold defTrajOfCirc_def,unfold stateSqLeq_def stateLeq_def, (rule allI)+, rule fSeqIsMono) prefer 2 show "defSqOfTrForm B \<sqsubseteq>sq defSqOfTrForm (B andT A)" proof(simp,unfold stateSqLeq_def stateLeq_def,(rule allI)+,rule lubSelf) qed qed ultimately show ?thesis by(auto dest:sqLeqIsTrans) qed qed qed consts andLists::"trajForm list =>trajForm" primrec andListsNil:"andLists []=chaos" andListsCons:"andLists (A#listA)= A andT (andLists listA)" lemma GateNodeFclosureValaux1: assumes a:"nl:netlists" shows "(ALL m. m mem inps --> isDefinedIn m nl)--> (pair0 mem zip inps (map (fclosure nl s) inps))-->pair0 ∈ rclosure nl s" (is "?P inps") proof(induct_tac inps,simp) fix a list assume a1:"?P list" show "?P (a#list)" (is "?P1 -->?P2-->?P3") proof(rule impI)+ assume b1:"?P1" and b2:"?P2" show "?P3" proof - from b2 have "pair0 =(a,fclosure nl s a) | pair0 mem zip list (map (fclosure nl s) list) " apply - apply simp apply(case_tac "(a, fclosure nl s a) = pair0") by auto moreover {assume c1:"pair0 =(a,fclosure nl s a)" from b1 c1 have "?P3" apply - apply simp apply(rule fclosureIsrclosure) by auto } moreover {assume c1:" pair0 mem zip list (map (fclosure nl s) list)" from b1 have c2:"(∀m. m mem list --> isDefinedIn m nl)" by auto from a1 this c1 have "?P3" by auto } ultimately show "?P3" by blast qed qed qed lemma GateNodeFclosureValaux1: assumes a1:"nl:netlists" and a2:"ALL m. m mem inps --> isDefinedIn m nl" and a3:"pair0 mem zip inps (map (fclosure nl s) inps)" shows "pair0 ∈ rclosure nl s" proof(cut_tac a1 a2 a3, blast dest:GateNodeFclosureValaux1) qed lemma GateNodeFclosureVal : assumes a1:"isClosed nl" and a3:"(Gate out inps tab):nl" and a4:"nl:netlists" and a5:"stateLs= map (fclosure nl s) inps" and a6:"∀l. l ∈ set tab --> length l = length inps" shows "fclosure nl s out=lub (funOfTab tab stateLs) (s out)" proof(rule fclosureGate) show "length stateLs = length inps" apply(cut_tac a5,simp) done next show "∀pair0. pair0 mem zip inps stateLs --> pair0 ∈ rclosure nl s" proof(rule allI,rule impI,cut_tac a5,simp) fix pair0 assume b1:"pair0 mem zip inps (map (fclosure nl s) inps)" show "pair0 ∈ rclosure nl s" proof(rule GateNodeFclosureValaux1) show "∀m. m mem inps --> isDefinedIn m nl" proof(rule allI,rule impI) fix m assume d1:" m mem inps " show "isDefinedIn m nl" proof(rule GateInNetlist) qed qed qed qed next show "nl ∈ netlists" by assumption next show "Gate out inps tab ∈ nl" by assumption qed (*lemma ftabIsMono: assumes a1:"ALL n. n mem inps --> s n \<sqsubseteq> t n" shows " funOfTab tab (map s inps) \<sqsubseteq> funOfTab tab (map t inps)" (is "?P tab") proof(induct_tac tab) show "?P []" by( simp add:leqReflexive) next fix a list0 assume b1:"?P list0" show "?P (a # list0)" apply - apply(simp (no_asm)) apply(rule orMono) apply(rule funOfLineIsMono) apply(rule_tac inps="inps" and nl="nl" and ?s1.0="s1" and ?s2.0="s2" in stateListLeq) apply(assumption)+ apply(cut_tac a1, blast) apply assumption apply(cut_tac a2, simp) thm funOfLineIsMono apply(cut_tac b1,simp) done qed*) lemma GateNodeVal : assumes a1:"isClosed nl" and a3:"(Gate n inps tab):nl" and a4:"nl:netlists" and a6:"∀l. l ∈ set tab --> length l = length inps" shows "funOfTab tab (map s inps) \<sqsubseteq> fclosure nl s n " proof - let ?stateLs="( map (fclosure nl s) inps)" have b1:" fclosure nl s n=lub (funOfTab tab ?stateLs) (s n)" apply(rule_tac inps="inps" in GateNodeFclosureVal) apply assumption+ apply simp by assumption then have b2:" (funOfTab tab (map (fclosure nl s) inps)) \<sqsubseteq>fclosure nl s n" proof(simp, rule_tac lubSelf) qed have b3:"(funOfTab tab (map s inps))\<sqsubseteq> (funOfTab tab (map (fclosure nl s) inps))" apply(cut_tac a6) apply(rule funOfTabIsMonoAny) apply(rule allI,rule impI) apply(rule fclosureIsExtensive1) by auto with b2 show ?thesis by(rule_tac leqIsTrans) qed lemma aux1:"T \<sqsubseteq> defSqOfTrForm (andLists (map Is1 (a # list))) 0 a" (is "?P list a") proof(case_tac list) assume b1:"list=[]" show "?P list a" proof(cut_tac b1,simp,rule lubSelf) qed next fix aa lista assume b1:"list = aa # lista" show "?P list a" proof(cut_tac b1,simp,case_tac "a=aa",auto intro: lubSelf) qed qed lemma aux2: assumes a1:"ALL n. s' n \<sqsubseteq> s n" shows "bv mem map s list -->( EX bv0. bv0\<sqsubseteq> bv& bv0 mem map s' list) " (is "?P list ") proof(induct_tac list) show "?P [] " proof(simp) qed next fix a list assume b1:"?P list" show "?P (a#list)" proof(rule impI ) assume c1:"bv mem map s (a # list)" from c1 have "bv =s a | bv mem map s list" apply - apply simp apply(case_tac "s a = bv") by auto moreover {assume d1:"bv =s a" have "(∃bv0. bv0 \<sqsubseteq> bv ∧ bv0 mem map s' (a # list))" apply(rule_tac x="s' a" in exI) apply(cut_tac a1 d1,auto) done } moreover {assume d1:" bv mem map s list" have "(∃bv0. bv0 \<sqsubseteq> bv ∧ bv0 mem map s' (a # list))" apply(cut_tac b1 d1,auto) done } ultimately show "(∃bv0. bv0 \<sqsubseteq> bv ∧ bv0 mem map s' (a # list))" by blast qed qed lemma aux3:"ALL n. defSqOfTrForm (andLists (map Is1 ( list))) 0 n \<sqsubseteq> defSqOfTrForm (andLists (map Is1 (a # list))) 0 n" (is "?P list") proof(case_tac list ) show "[|list = []|]==>?P list" apply (simp add:XIsLeastValue) done next fix aa lista assume b1:"list = aa # lista" show " ?P list" proof fix n show "defSqOfTrForm (andLists (map Is1 list)) 0 n \<sqsubseteq> defSqOfTrForm (andLists (map Is1 (a # list))) 0 n" (is "?LHS \<sqsubseteq>?RHS") proof - have b2:" ?RHS=(if (n = a )then lub T (?LHS) else lub X (?LHS))" proof(cut_tac b1,simp) qed have b3:"?LHS\<sqsubseteq> (if (n = a )then lub T (?LHS) else lub X (?LHS))" apply(case_tac "n=a") apply simp apply (subgoal_tac "lub T (defSqOfTrForm (andLists (map Is1 list)) 0 a) = lub (defSqOfTrForm (andLists (map Is1 list)) 0 a) T") apply(simp,rule lubSelf) apply(simp add:lubComm) apply simp apply (subgoal_tac "lub X (defSqOfTrForm (andLists (map Is1 list)) 0 n) = lub (defSqOfTrForm (andLists (map Is1 list)) 0 n) X") apply(simp,rule lubSelf) apply(simp add:lubComm) done from b2 b3 show "?LHS\<sqsubseteq>?RHS" by auto qed qed qed lemma aux2: assumes a1:"ALL n. s' n \<sqsubseteq> s n" and a2:"bv mem map s list" shows "( EX bv0. bv0\<sqsubseteq> bv& bv0 mem map s' list) " proof(cut_tac a1 a2, blast dest:aux2) qed lemma andTabPropT: assumes a1:"isAndTab tab" and a3:"Gate out inps tab:nl" and a4:"nl:netlists" and a5:"isClosed nl" and a6:" ∀ l. (l∈ (set tab)) --> length l=length inps" shows "cktSat nl ((andLists (map (λn. Is1 n) inps) ) \<leadsto> (Is1 out))" proof(rule fundmental1) show "defSqOfTrForm (Is1 out) \<sqsubseteq>sq defTrajOfCirc (andLists (map Is1 inps)) nl" proof( unfold defTrajOfCirc_def, simp, unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n show "(if t = 0 ∧ n = out then T else X) \<sqsubseteq> fSeq nl (defSqOfTrForm (andLists (map Is1 inps))) t n" (is "?LHS\<sqsubseteq>?RHS") proof - have "(t = 0 ∧ n = out)| ~(t = 0 ∧ n = out)" by auto moreover {assume b1:"(t = 0 ∧ n = out)" have b2:"?LHS =T" by(cut_tac b1,simp) have b3:"T\<sqsubseteq> ?RHS" proof(cut_tac b1,simp) let ?s=" (defSqOfTrForm (andLists (map Is1 inps)) 0) " have c1:"funOfTab tab (map ?s inps) \<sqsubseteq> fclosure nl ?s out" by(rule GateNodeVal) have c2:"∀bv. bv mem map (defSqOfTrForm (andLists (map Is1 inps)) 0) inps --> T \<sqsubseteq> bv" (is "?P inps") proof(induct_tac inps) show "?P []" by simp next fix a list assume d1:"?P list" show "?P (a#list)" (is "ALL bv. ?P1 bv --> ?P2 bv") proof(rule allI,rule impI) fix bv assume e1:"?P1 bv" show "?P2 bv" proof - let ?s="defSqOfTrForm (andLists (map Is1 (a # list))) 0" from e1 have f1:"bv=?s a | bv mem (map ?s list)" apply - apply(simp del:andListsNil andListsCons) apply(case_tac "defSqOfTrForm (andLists (Is1 a # map Is1 list)) 0 a = bv") apply simp apply simp done moreover {assume f1:"bv=?s a" have f0:"T \<sqsubseteq>?s a" by(rule aux1) from f1 this have "T \<sqsubseteq> bv" by simp } moreover { assume f1:"bv mem (map ?s list)" let ?t="defSqOfTrForm (andLists (map Is1 list)) 0" have "list=[] |list ~=[]" by auto moreover {assume g1:"list=[]" from f1 g1 have "T \<sqsubseteq> bv" by simp } moreover {assume g1:"list~=[]" have "( EX bv0. bv0\<sqsubseteq> bv& bv0 mem map ?t list)" (is "EX bv0. ?P bv0") proof(rule_tac s="?s" in aux2) show "∀n. ?t n \<sqsubseteq> ?s n" proof(cut_tac g1,rule aux3) qed qed then obtain bv0 where g2:" ?P bv0" by blast from d1 g2 have "T \<sqsubseteq> bv0" by blast with g2 have "T \<sqsubseteq> bv" by (blast dest:leqIsTrans) } ultimately have "T \<sqsubseteq> bv" by blast } ultimately show "T \<sqsubseteq> bv" by blast qed qed qed then have c2:"T\<sqsubseteq> funOfTab tab (map ?s inps)" proof(rule andTabPropT) show "length (map (defSqOfTrForm (andLists (map Is1 inps)) 0) inps) = length (hd tab)" (is "?LHS=?RHS") proof - have "?LHS=length inps" by simp also have "…=?RHS" proof(rule sym) have "EX y ys. tab=y#ys" (is "EX y ys. ?P y ys") proof(cut_tac a1, unfold isAndTab_def, simp add: length_Suc_conv,auto ) qed then obtain y and ys where d1:"?P y ys" by blast then have d2:"hd tab =y" by simp from d1 have d3:"y:set tab" by simp with d2 show "length (hd tab) = length inps" apply - apply(cut_tac a6) apply(drule_tac x="y" in spec) apply auto done qed finally show "?LHS=?RHS" by auto qed qed from c1 c2 show "T \<sqsubseteq> fclosure nl (defSqOfTrForm (andLists (map Is1 inps)) 0) out" thm leqIsTrans apply (rule_tac leqIsTrans) by auto qed with b2 have "?LHS\<sqsubseteq>?RHS" by simp } moreover {assume b1:"~(t = 0 ∧ n = out)" from b1 have "?LHS =X" by auto then have "?LHS \<sqsubseteq> ?RHS" apply - apply(simp) apply(rule XIsLeastValue) done } ultimately show "?LHS \<sqsubseteq> ?RHS" by blast qed qed qed lemma andTabPropF: assumes a1:"isAndTab tab" and a3:"Gate out inps tab:nl" and a4:"nl:netlists" and a5:"isClosed nl" and a6:" ∀ l. (l∈ (set tab)) --> length l=length inps" and a7:"inpsi mem inps" shows "cktSat nl (( Is0 inpsi) \<leadsto> (Is0 out))" proof(rule fundmental1) show "defSqOfTrForm (Is0 out) \<sqsubseteq>sq defTrajOfCirc ( Is0 inpsi) nl" proof( unfold defTrajOfCirc_def, simp, unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n show "(if t = 0 ∧ n = out then F else X) \<sqsubseteq> fSeq nl (λt m. if t = 0 ∧ m = inpsi then F else X) t n" (is "?LHS\<sqsubseteq>?RHS") proof - have "(t = 0 ∧ n = out)| ~(t = 0 ∧ n = out)" by auto moreover {assume b1:"(t = 0 ∧ n = out)" have b2:"?LHS =F" by(cut_tac b1,simp) have b3:"F\<sqsubseteq> ?RHS" proof(cut_tac b1,simp) let ?s=" (λm. if m = inpsi then F else X) " let ?bv="?s inpsi" have c1:"funOfTab tab (map ?s inps) \<sqsubseteq> fclosure nl ?s out" by(rule GateNodeVal) thm GateNodeVal have c2:" F \<sqsubseteq> ?bv" proof(simp,rule leqReflexive) qed have c2a:" ?bv mem (map ?s inps)" proof - have d1:"inpsi:set inps" apply(cut_tac a7) apply( simp add:mem_iff) done then have "?s inpsi:?s `(set inps)" by(rule_tac imageI) then have "(if inpsi = inpsi then F else X):set (map ?s inps)" apply(subgoal_tac "set (map ?s inps)=?s `(set inps)") apply force thm set_map apply(simp (no_asm_use) add: set_map) done then show ?thesis thm mem_iff apply - thm mem_iff[THEN sym] apply(simp add: mem_iff) done qed with c2 have c2:"F\<sqsubseteq> funOfTab tab (map ?s inps)" proof(rule andTabPropF) show "length (map ?s inps) = length (hd tab)" (is "?LHS=?RHS") proof - have "?LHS=length inps" by simp also have "…=?RHS" proof(rule sym) have "EX y ys. tab=y#ys" (is "EX y ys. ?P y ys") proof(cut_tac a1, unfold isAndTab_def, simp add: length_Suc_conv,auto ) qed then obtain y and ys where d1:"?P y ys" by blast then have d2:"hd tab =y" by simp from d1 have d3:"y:set tab" by simp with d2 show "length (hd tab) = length inps" apply - apply(cut_tac a6) apply(drule_tac x="y" in spec) apply auto done qed finally show "?LHS=?RHS" by auto qed qed from c1 c2 show "F \<sqsubseteq> fclosure nl ?s out" thm leqIsTrans apply (rule_tac leqIsTrans) by auto qed with b2 have "?LHS\<sqsubseteq>?RHS" by simp } moreover {assume b1:"~(t = 0 ∧ n = out)" from b1 have "?LHS =X" by auto then have "?LHS \<sqsubseteq> ?RHS" apply - apply(simp) apply(rule XIsLeastValue) done } ultimately show "?LHS \<sqsubseteq> ?RHS" by blast qed qed qed consts PosAssertOfLine::"node list => LINE =>trajForm list" primrec PosAss1: " PosAssertOfLine inps [] =[]" PosAss2: " PosAssertOfLine (inps) (a#line)= (let otherAss=PosAssertOfLine (tl inps) ( line) in (case (a) of ONE => (Is1 (hd inps))#otherAss| ZERO => (Is0 (hd inps))#otherAss| DontCare => chaos#otherAss))" (*lemma aux4: assumes a1:"(PosAssertOfLine inps list) ~=[] " shows "listLeq (map (defSqOfTrForm (andLists (PosAssertOfLine inps ( list))) 0) zs) (map (defSqOfTrForm (andLists (PosAssertOfLine inps (a # list))) 0) zs)" proof(cut_tac a1, simp) constdefs PosAssertOfLineTab::"node list =>PLA =>trajForm list" " PosAssertOfLineTab inps tab==map PosAssertOfLine tab"*) lemma listLeqstate:"listLeq (map s zs) (map (λu. lub (if u = z then v else X) (s u)) zs)" (is "?P zs") proof(induct_tac zs) show "?P []" by auto next fix a list assume b1:"?P list" show "?P (a#list)" proof(cut_tac b1,simp,case_tac "a=z",auto) show "s z \<sqsubseteq> lub v (s z)" proof - have "lub v (s z)=lub (s z) v" by(auto intro:lubComm) then show "s z \<sqsubseteq> lub v (s z)" by (auto intro:lubSelf) qed next show "s a \<sqsubseteq> lub X (s a)" proof - have "lub X (s a)=lub (s a) X" by(auto intro:lubComm) then show "s a \<sqsubseteq> lub X (s a)" by (auto intro:lubSelf) qed qed qed (*lemma listLeqstate1:"listLeq (map s zs) (map (λu. lub (if u = z then v else X) (s u)) zs)" (is "?P zs") proof(induct_tac zs) show "?P []" by auto next fix a list assume b1:"?P list" show "?P (a#list)" proof(cut_tac b1,simp,case_tac "a=z",auto) show "s z \<sqsubseteq> lub v (s z)" proof - have "lub v (s z)=lub (s z) v" by(auto intro:lubComm) then show "s z \<sqsubseteq> lub v (s z)" by (auto intro:lubSelf) qed next show "s a \<sqsubseteq> lub X (s a)" proof - have "lub X (s a)=lub (s a) X" by(auto intro:lubComm) then show "s a \<sqsubseteq> lub X (s a)" by (auto intro:lubSelf) qed qed qed *) lemma PosAss2PosVal: shows "ALL inps. length line=length inps --> listLeq ( PosValOfLine line) ( map (defSqOfTrForm (andLists (PosAssertOfLine inps line)) 0 ) inps) " (is "?P line") proof(induct_tac line) show "?P []" proof(simp) qed next fix a list assume b1:"?P list" show "?P (a#list)" proof(rule allI,rule impI)+ fix inps assume c1:"length (a # list) = length (inps:: node list)" show "listLeq ( PosValOfLine (a # list)) ( map (defSqOfTrForm (andLists (PosAssertOfLine inps (a # list))) 0 ) inps) " (is "?R (a # list) inps") proof(case_tac "a::LIT") assume d1:"a = ONE " show "?R (a # list) inps" proof (cut_tac d1) have " length inps =Suc (length list) " proof(cut_tac c1, auto) qed then have b2:"EX y ys.( inps = y # ys & length ys = length list)" by (simp add:length_Suc_conv) then obtain z zs where b2:"inps = z # zs& length zs = length list" by auto have e0:"(PosValOfLine (a # list))= T # PosValOfLine list" proof(cut_tac d1, simp add:Let_def) qed have e1:" (PosAssertOfLine inps (a # list))= Is1 (z) # PosAssertOfLine (zs) list " proof(cut_tac b2 d1,simp add:Let_def) qed then have e1a:"andLists (PosAssertOfLine inps (a # list))=andLists ( Is1 z # PosAssertOfLine (zs) list)" by simp show ?thesis proof(cut_tac b2 d1, simp add:Let_def,rule conjI,rule lubSelf) let ?zs="map (λu. lub (if u = z then T else X) (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 u)) zs" let ?ys="(map (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0) zs)" show "listLeq (PosValOfLine list) (map (λu. lub (if u = z then T else X) (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 u)) zs)" proof - have h1:"listLeq (PosValOfLine list) ?ys" proof(cut_tac b1 b2,drule_tac x="zs" in spec,simp) qed have h2:"listLeq ?ys ?zs" thm listLeqstate proof(rule listLeqstate) qed with h1 show " listLeq (PosValOfLine list) ?zs" by(rule_tac listLeqTrans) qed qed qed next assume d1:"a = ZERO " show "?R (a # list) inps" proof (cut_tac d1) have " length inps =Suc (length list) " proof(cut_tac c1, auto) qed then have b2:"EX y ys.( inps = y # ys & length ys = length list)" by (simp add:length_Suc_conv) then obtain z zs where b2:"inps = z # zs& length zs = length list" by auto have e0:"(PosValOfLine (a # list))= F # PosValOfLine list" proof(cut_tac d1, simp add:Let_def) qed have e1:" (PosAssertOfLine inps (a # list))= Is0 (z) # PosAssertOfLine (zs) list " proof(cut_tac b2 d1,simp add:Let_def) qed then have e1a:"andLists (PosAssertOfLine inps (a # list))=andLists ( Is0 z # PosAssertOfLine (zs) list)" by simp show ?thesis proof(cut_tac b2 d1, simp add:Let_def,rule conjI,rule lubSelf) let ?zs="map (λu. lub (if u = z then F else X) (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 u)) zs" let ?ys="(map (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0) zs)" show "listLeq (PosValOfLine list) (map (λu. lub (if u = z then F else X) (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 u)) zs)" proof - have h1:"listLeq (PosValOfLine list) ?ys" proof(cut_tac b1 b2,drule_tac x="zs" in spec,simp) qed have h2:"listLeq ?ys ?zs" thm listLeqstate proof(rule listLeqstate) qed with h1 show " listLeq (PosValOfLine list) ?zs" by(rule_tac listLeqTrans) qed qed qed next assume d1:"a = DontCare " show "?R (a # list) inps" proof (cut_tac d1) have " length inps =Suc (length list) " proof(cut_tac c1, auto) qed then have b2:"EX y ys.( inps = y # ys & length ys = length list)" by (simp add:length_Suc_conv) then obtain z zs where b2:"inps = z # zs& length zs = length list" by auto have e0:"(PosValOfLine (a # list))= X # PosValOfLine list" proof(cut_tac d1, simp add:Let_def) qed have e1:" (PosAssertOfLine inps (a # list))= chaos # PosAssertOfLine (zs) list " proof(cut_tac b2 d1,simp add:Let_def) qed then have e1a:"andLists (PosAssertOfLine inps (a # list))=andLists ( chaos # PosAssertOfLine (zs) list)" by simp show ?thesis proof(cut_tac b2 d1, simp add:Let_def,rule conjI,rule lubSelf) let ?zs="map (λu. lub X (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 u)) zs" let ?ys="(map (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0) zs)" show "listLeq (PosValOfLine list) (map (λu. lub X (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 u)) zs)" proof - have h1:"listLeq (PosValOfLine list) ?ys" proof(cut_tac b1 b2,drule_tac x="zs" in spec,simp) qed have h2:"?ys= ?zs" thm listLeqstate proof(simp,rule ballI ) fix x have "lub X (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 x)= lub (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 x) X" by(rule lubComm) then show "defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 x = lub X (defSqOfTrForm (andLists (PosAssertOfLine zs list)) 0 x)" proof(simp add:XIsLubUnit) qed qed with h1 show " listLeq (PosValOfLine list) ?zs" by simp qed qed qed qed qed qed lemma PosAss2PosVal: "[| length line=length inps |]==> listLeq ( PosValOfLine line) ( map (defSqOfTrForm (andLists (PosAssertOfLine inps line)) 0 ) inps) " apply(cut_tac PosAss2PosVal[where line="line"]) apply blast done lemma orTabPropT: assumes a3:"Gate out inps tab:nl" and a4:"nl:netlists" and a5:"isClosed nl" and a6:" ∀ l. (l∈ (set tab)) --> length l=length inps" and a7:"line mem tab" shows "cktSat nl ( andLists ( PosAssertOfLine inps line) \<leadsto> (Is1 out))" proof(rule fundmental1) show "defSqOfTrForm (Is1 out) \<sqsubseteq>sq defTrajOfCirc (andLists (PosAssertOfLine inps line)) nl" proof( unfold defTrajOfCirc_def, simp, unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n show "(if t = 0 ∧ n = out then T else X) \<sqsubseteq> fSeq nl (defSqOfTrForm (andLists (PosAssertOfLine inps line))) t n" (is "?LHS\<sqsubseteq>?RHS") proof - have "(t = 0 ∧ n = out)| ~(t = 0 ∧ n = out)" by auto moreover {assume b1:"(t = 0 ∧ n = out)" have b2:"?LHS =T" by(cut_tac b1,simp) have b3:"T\<sqsubseteq> ?RHS" proof(cut_tac b1,simp) let ?s=" (defSqOfTrForm (andLists (PosAssertOfLine inps line)) 0) " let ?bv="?s inpsi" have c1:"funOfTab tab (map ?s inps) \<sqsubseteq> fclosure nl ?s out" by(rule GateNodeVal) thm GateNodeVal have c2:"T\<sqsubseteq> funOfTab tab (map ?s inps)" proof(cut_tac a7 a6,rule orTabPropT,auto, drule_tac x="line" in spec,simp add:mem_iff, rule PosAss2PosVal,simp add:mem_iff ) qed with c1 show "T \<sqsubseteq> fclosure nl (defSqOfTrForm (andLists (PosAssertOfLine inps line)) 0) out" by(rule_tac leqIsTrans) qed from b2 this have "?LHS\<sqsubseteq>?RHS" by simp } moreover {assume b1:"~(t = 0 ∧ n = out)" have "?LHS\<sqsubseteq>?RHS" proof(cut_tac b1,auto intro!:XIsLeastValue)qed } ultimately show "?LHS\<sqsubseteq>?RHS" by blast qed qed qed consts isNegAssertOfLine::" trajForm=> node list => LINE => bool" primrec "isNegAssertOfLine (Is1 n) inps line = ( n mem inps & (EX pair. pair mem zip inps line & (fst pair=n) & (snd pair)=ZERO))" "isNegAssertOfLine (Is0 n) inps line = ( n mem inps & (EX pair. pair mem zip inps line & (fst pair=n) & (snd pair)=ONE))" "isNegAssertOfLine (f1 andT f2) inps line = False" "isNegAssertOfLine (P ->T f2) inps line = False" "isNegAssertOfLine (Next f2) inps line = False" "isNegAssertOfLine (chaos) inps line = False" lemma aux1:"ALL line. pair mem (zip inps line )--> ((snd pair), f (fst pair)) mem (zip line (map f inps))" (is "?P inps line ") proof(induct_tac inps ) show "?P [] line" by simp next fix a list assume b1:"?P list line" show "?P (a#list) line" proof(rule allI,rule impI) fix line assume c1:"pair mem zip (a # list) line" show "(snd pair, f (fst pair)) mem zip line (map f (a # list))" proof (case_tac "line::'a list") assume d1:"line = []" from c1 d1 have False by simp then show ?thesis by simp next fix a0 list0 assume d1:"line =a0#list0" from c1 d1 have d2:"pair=(a,a0) |pair mem zip ( list) list0" apply - apply simp apply(case_tac "(a, a0) = pair") by auto moreover {assume e1:"pair=(a,a0)" have ?thesis by(cut_tac e1 d1,auto) } moreover {assume e1:"pair mem zip ( list) list0" from b1 e1 have e2:"(snd pair, f (fst pair)) mem zip list0 (map f list)" by blast with d1 have ?thesis by simp } ultimately show ?thesis by blast qed qed qed lemma aux1:" pair mem (zip inps line )==> ((snd pair), f (fst pair)) mem (zip line (map f inps))" apply(cut_tac aux1[where inps="inps" and pair="pair" and f="f"],blast) done lemma aux2:"[|pair mem (zip inps line ); (snd pair) = ZERO; fst pair = n; contradict ((snd pair), f (fst pair)) |]==> T \<sqsubseteq> f n" apply(drule_tac f="f" in aux1) apply(simp add:contradict_def) done lemma aux2: assumes a1:"ALL n. s' n \<sqsubseteq> s n" shows "ALL line. length line=length inps-->asOfVal mem zip line (map s' inps) & contradict asOfVal -->(EX asOfVal. asOfVal mem zip line (map s inps) & contradict asOfVal)" (is "?P inps ") proof(induct_tac inps) show "?P [] " proof(simp) qed next fix a list assume b1:"?P list" show "?P (a#list)" proof(rule allI,(rule impI )+) fix line assume c1:" asOfVal mem zip line (map s' (a # list)) ∧ contradict asOfVal " and c2:"length line =length (a#list)" have " length line =Suc (length list) " proof(cut_tac c2, auto) qed then have b2:"EX y ys.( line = y # ys & length ys = length list)" by (simp add:length_Suc_conv) then obtain z zs where b2:"line = z # zs& length zs = length list" by auto with c1 have "asOfVal =(z,s' a) | asOfVal mem zip zs (map s' list)" apply - apply simp apply(case_tac "(z, s' a) = asOfVal",auto) done moreover {assume d1:"asOfVal =(z,s' a)" let ?asOfVal="(z,s a)" have d2:"?asOfVal mem zip line (map s (a # list)) ∧ contradict ?asOfVal" apply(cut_tac c1 d1 b2,simp add:contradict_def) apply(cut_tac a1) apply(auto dest:leqIsTrans) done then have "∃asOfVal. asOfVal mem zip line (map s (a # list)) ∧ contradict asOfVal" by blast } moreover {assume d1:"asOfVal mem zip zs (map s' list)" have " (∃asOfVal. asOfVal mem zip zs (map s list) ∧ contradict asOfVal)" (is "EX asOfVal. ?R zs list asOfVal") apply(cut_tac b1 b2 d1 c1) by blast then obtain asOfVal where e1:" ?R zs list asOfVal" by blast with b2 have "?R line (a#list) asOfVal " by simp then have "EX asOfVal. ?R line (a#list) asOfVal " by blast } ultimately show "∃asOfVal. asOfVal mem zip line (map s (a # list)) ∧ contradict asOfVal" by blast qed qed lemma aux2: assumes a1:"ALL n. s' n \<sqsubseteq> s n" and a2:"length line=length inps" and a3:"asOfVal mem zip line (map s' inps) " and a4:"contradict asOfVal" shows "(EX asOfVal. asOfVal mem zip line (map s inps) & contradict asOfVal)" proof(cut_tac a1 a2 a3 a4 aux2[where inps="inps" and s="s" and s'="s'" and asOfVal="asOfVal"],auto) qed lemma aux2: assumes a1:"asOfVal mem zip line (map (defSqOfTrForm (andLists list) 0) inps) ∧ contradict asOfVal" and a2:"length line=length inps" shows "∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict asOfVal" proof(rule_tac s'="defSqOfTrForm (andLists ( list)) 0 " and asOfVal="asOfVal" in aux2) show "∀n. defSqOfTrForm (andLists list) 0 n \<sqsubseteq> defSqOfTrForm (andLists (aa # list)) 0 n" apply(rule allI,simp) apply(subgoal_tac " lub (defSqOfTrForm aa 0 n) (defSqOfTrForm (andLists list) 0 n)= lub (defSqOfTrForm (andLists list) 0 n)(defSqOfTrForm aa 0 n)" ) apply simp apply(rule lubSelf) apply(rule lubComm) done prefer 2 show "asOfVal mem zip line (map (defSqOfTrForm (andLists list) 0) inps)" by(cut_tac a1,simp) prefer 2 show "contradict asOfVal" by(cut_tac a1,simp) qed (*lemma aux2:"ALL line. pair mem zip inps line --> fst pair = n --> snd pair = ZERO --> asOfVal mem zip line (map f inps) --> contradict asOfVal --> T \<sqsubseteq> f n"(is "?P inps line ") proof(induct_tac inps ) show "?P [] line" by simp next fix a list assume b1:"?P list line" show "?P (a#list) line" proof(rule allI,(rule impI)+) fix line assume c1:"pair mem zip (a # list) line" and c2:"fst pair = n" and c3:"snd pair = ZERO" and c4:"asOfVal mem zip line (map f (a # list))" and c5:"contradict asOfVal" show " T \<sqsubseteq> f n" proof (case_tac "line::sign list") assume d1:"line = []" from c1 d1 have False by simp then show ?thesis by simp next fix a0 list0 assume d1:"line =a0#list0" from c1 d1 have d2:"pair=(a,a0) |pair mem zip ( list) list0" apply - apply simp apply(case_tac "(a, a0) = pair") by auto moreover {assume e1:"pair=(a,a0)" have ?thesis apply(cut_tac e1 d1 c2 c3 c4 c5) apply simp } moreover {assume e1:"pair mem zip ( list) list0" from b1 e1 have e2:"(snd pair, f (fst pair)) mem zip list0 (map f list)" by blast with d1 have ?thesis by simp } ultimately show ?thesis by blast qed qed qed*) lemma negAux1: "(Is1 n) mem asLists --> n mem inps -->length line=length inps--> (∃pair. pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ZERO)--> (EX asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)" (is "?P asLists") proof(induct_tac asLists) show "?P []" by simp next fix aa list assume b1:"?P list" show "?P (aa#list)" proof(rule impI)+ assume c1:"Is1 n mem aa # list" and c3:"n mem inps" and c4:"length line=length inps" and c2:"∃pair. pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ZERO" show "∃asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict asOfVal" (is "EX asOfVal. ?Q asOfVal") proof - let ?s="(defSqOfTrForm (andLists (aa # list)) 0)" let ?v="?s n" from c2 obtain pair where d0:"pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ZERO" by blast from c1 have d1:"aa=Is1 n | (Is1 n) mem list" apply - apply simp apply(case_tac "aa = Is1 n",auto) done moreover {assume d1:"aa=Is1 n" have "EX asOfVal. ?Q asOfVal" proof(rule_tac x="(ZERO, ?v)" in exI) from d0 have e1:"(ZERO, defSqOfTrForm (andLists (aa # list)) 0 n) mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps)" thm aux1 apply - apply(erule conjE)+ apply(drule_tac f="?s" in aux1) by auto have e2:" contradict (ZERO, defSqOfTrForm (andLists (aa # list)) 0 n)" proof(cut_tac d1,simp add:contradict_def,rule lubSelf) qed with e1 show "(ZERO, defSqOfTrForm (andLists (aa # list)) 0 n) mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict (ZERO, defSqOfTrForm (andLists (aa # list)) 0 n)" by simp qed } moreover {assume d1:"(Is1 n) mem list" from d1 b1 c3 c2 c4 have e1:"(∃asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists list) 0) inps) ∧ contradict asOfVal)" (is "∃asOfVal. ?R asOfVal list") by blast then obtain asOfVal where e1:"?R asOfVal list" by blast with c4 have " ∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict asOfVal" apply - apply(rule aux2) by auto } ultimately show ?thesis by blast qed qed qed lemma negAux1: "[|(Is1 n) mem asLists ; n mem inps ;length line=length inps; (∃pair. pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ZERO)|]==> (EX asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)" apply(cut_tac negAux1,auto) done lemma negAux2: "(Is0 n) mem asLists --> n mem inps -->length line=length inps--> (∃pair. pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ONE)--> (EX asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)" (is "?P asLists") proof(induct_tac asLists) show "?P []" by simp next fix aa list assume b1:"?P list" show "?P (aa#list)" proof(rule impI)+ assume c1:"Is0 n mem aa # list" and c3:"n mem inps" and c4:"length line=length inps" and c2:"∃pair. pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ONE" show "∃asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict asOfVal" (is "EX asOfVal. ?Q asOfVal") proof - let ?s="(defSqOfTrForm (andLists (aa # list)) 0)" let ?v="?s n" from c2 obtain pair where d0:"pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ONE" by blast from c1 have d1:"aa=Is0 n | (Is0 n) mem list" apply - apply simp apply(case_tac "aa = Is0 n",auto) done moreover {assume d1:"aa=Is0 n" have "EX asOfVal. ?Q asOfVal" proof(rule_tac x="(ONE, ?v)" in exI) from d0 have e1:"(ONE, defSqOfTrForm (andLists (aa # list)) 0 n) mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps)" thm aux1 apply - apply(erule conjE)+ apply(drule_tac f="?s" in aux1) by auto have e2:" contradict (ONE, defSqOfTrForm (andLists (aa # list)) 0 n)" proof(cut_tac d1,simp add:contradict_def,rule lubSelf) qed with e1 show "(ONE, defSqOfTrForm (andLists (aa # list)) 0 n) mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict (ONE, defSqOfTrForm (andLists (aa # list)) 0 n)" by simp qed } moreover {assume d1:"(Is0 n) mem list" from d1 b1 c3 c2 c4 have e1:"(∃asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists list) 0) inps) ∧ contradict asOfVal)" (is "∃asOfVal. ?R asOfVal list") by blast then obtain asOfVal where e1:"?R asOfVal list" by blast with c4 have " ∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict asOfVal" apply - apply(rule aux2) by auto } ultimately show ?thesis by blast qed qed qed lemma negAux2: "[|(Is0 n) mem asLists ; n mem inps ;length line=length inps; (∃pair. pair mem zip inps ( line) ∧ fst pair = n ∧ snd pair = ONE)|]==> (EX asOfVal. asOfVal mem zip ( line) (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)" apply(cut_tac negAux2,auto) done lemma negAss2NegVal: shows " ALL inps. (∃as. as mem asLists ∧ isNegAssertOfLine as inps line)--> length line=length inps --> (∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)" (is "?P line") proof(induct_tac line) show "?P []" proof(rule allI, rule impI,erule_tac exE,case_tac "as::trajForm",auto) qed next fix a line assume b1:"?P line" show "?P (a#line)" proof(rule allI, (rule impI)+,erule_tac exE ) fix inps as assume c1:"as mem asLists ∧ isNegAssertOfLine as inps (a # line) " and c2:"length (a # line) = length inps" from c1 show "∃asOfVal. asOfVal mem zip (a # line) (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal" proof(case_tac "as::trajForm") fix nat assume d1:"as mem asLists ∧ isNegAssertOfLine as inps (a # line)" and d2:"as = Is1 nat" show "∃asOfVal. asOfVal mem zip (a # line) (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal" proof(cut_tac d1 d2 c2,simp_all del:split_paired_Ex,(erule conjE)+,rule negAux1,auto)qed next fix nat assume d1:"as mem asLists ∧ isNegAssertOfLine as inps (a # line)" and d2:"as = Is0 nat" show "∃asOfVal. asOfVal mem zip (a # line) (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal" proof(cut_tac d1 d2 c2,simp_all del:split_paired_Ex,(erule conjE)+,rule negAux2,auto)qed qed( auto) qed qed lemma negAss2NegVal: shows "[| (∃as. as mem asLists ∧ isNegAssertOfLine as inps line); length line=length inps|] ==> (∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)" apply(cut_tac negAss2NegVal[where asLists="asLists"]) apply blast done lemma orTabPropF: assumes a3:"Gate out inps tab:nl" and a4:"nl:netlists" and a5:"isClosed nl" and a6:" ∀ l. (l∈ (set tab)) --> length l=length inps" and a7:"ALL line. line mem tab --> (EX as. as mem asLists & isNegAssertOfLine as inps line)" shows "cktSat nl ( andLists ( asLists) \<leadsto> (Is0 out))" proof(rule fundmental1) show "defSqOfTrForm (Is0 out) \<sqsubseteq>sq defTrajOfCirc (andLists (asLists)) nl" proof( unfold defTrajOfCirc_def, simp, unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n show "(if t = 0 ∧ n = out then F else X) \<sqsubseteq> fSeq nl (defSqOfTrForm (andLists asLists)) t n" (is "?LHS\<sqsubseteq>?RHS") proof - have "(t = 0 ∧ n = out)| ~(t = 0 ∧ n = out)" by auto moreover {assume b1:"(t = 0 ∧ n = out)" have b2:"?LHS =F" by(cut_tac b1,simp) have b3:"F\<sqsubseteq> ?RHS" proof(cut_tac b1,simp) let ?s=" (defSqOfTrForm (andLists asLists) 0) " have c1:"funOfTab tab (map ?s inps) \<sqsubseteq> fclosure nl ?s out" by(rule GateNodeVal) thm GateNodeVal have c2:"F\<sqsubseteq> funOfTab tab (map ?s inps)" proof(cut_tac a7 a6,rule orTabPropF1,simp, rule allI,rule impI, subgoal_tac "line ∈ set tab",rule negAss2NegVal,blast,blast,simp add:mem_iff) qed with c1 show "F \<sqsubseteq> fclosure nl (defSqOfTrForm (andLists (asLists)) 0) out" by(rule_tac leqIsTrans) qed from b2 this have "?LHS\<sqsubseteq>?RHS" by simp } moreover {assume b1:"~(t = 0 ∧ n = out)" have "?LHS\<sqsubseteq>?RHS" proof(cut_tac b1,auto intro!:XIsLeastValue)qed } ultimately show "?LHS\<sqsubseteq>?RHS" by blast qed qed qed lemma steRef: assumes a1:"nl:netlists" and a2:"isClosed nl" shows " cktSat nl (A \<leadsto> A)" proof(rule fundmental1,assumption) show "defSqOfTrForm A \<sqsubseteq>sq defTrajOfCirc A nl" thm fSeqIsMono proof(unfold defTrajOfCirc_def,rule fSeqIsExtensive)qed qed lemma steConjI: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"cktSat nl (A\<leadsto> B)" and a4:"cktSat nl (A\<leadsto> C)" shows " cktSat nl (A \<leadsto> B andT C)" proof (rule fundmental1,assumption) show "defSqOfTrForm (B andT C) \<sqsubseteq>sq defTrajOfCirc A nl" proof(unfold stateSqLeq_def stateLeq_def,(rule allI)+) fix t n show "defSqOfTrForm (B andT C) t n \<sqsubseteq> defTrajOfCirc A nl t n" proof(simp,rule lubMeans) thm fundmental2 from a3 have "defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A nl " proof(rule_tac fundmental2) qed then show "defSqOfTrForm B t n \<sqsubseteq> defTrajOfCirc A nl t n" by(unfold stateSqLeq_def stateLeq_def, blast) next from a3 have "defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc A nl " proof(rule_tac fundmental2) qed then show "defSqOfTrForm C t n \<sqsubseteq> defTrajOfCirc A nl t n" by(unfold stateSqLeq_def stateLeq_def, blast) qed qed qed lemma steTrans: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"cktSat nl (A\<leadsto> B)" and a4:"cktSat nl (B\<leadsto> C)" shows " cktSat nl (A \<leadsto> C)" proof (rule fundmental1,assumption) show "defSqOfTrForm (C) \<sqsubseteq>sq defTrajOfCirc A nl" proof - have "defSqOfTrForm (C) \<sqsubseteq>sq defTrajOfCirc B nl" thm fundmental2 proof(rule_tac fundmental2) qed also have "…\<sqsubseteq>sq defTrajOfCirc A nl" proof - have "defSqOfTrForm (B) \<sqsubseteq>sq defTrajOfCirc A nl" proof(rule_tac fundmental2) qed then have b1:"fSeq nl (defSqOfTrForm B) \<sqsubseteq>sq fSeq nl (defTrajOfCirc A nl)" proof(unfold stateSqLeq_def stateLeq_def, (rule_tac allI)+,rule_tac fSeqIsMono) qed have b2:"defTrajOfCirc B nl =fSeq nl (defSqOfTrForm B)" apply(unfold defTrajOfCirc_def) by simp with b1 b2 have b3:"defTrajOfCirc B nl \<sqsubseteq>sq fSeq nl (defTrajOfCirc A nl)" by simp have b4:" fSeq nl (defTrajOfCirc A nl) = (defTrajOfCirc A nl)" proof - have " (defTrajOfCirc A nl):trajOfCirc nl" proof(rule defTrajIsTraj) qed then show "fSeq nl (defTrajOfCirc A nl) = (defTrajOfCirc A nl)" proof(unfold trajOfCirc_def,auto) qed qed with b3 show "defTrajOfCirc B nl \<sqsubseteq>sq (defTrajOfCirc A nl)" by simp qed ultimately show ?thesis by(blast dest:sqLeqIsTrans) qed qed lemma steImpI: assumes a3:"g1 --> (cktSat nl (A\<leadsto> B))" and a1:"nl:netlists" and a2:"isClosed nl" shows " cktSat nl (A \<leadsto> g1->T B)" proof(case_tac g1) assume b1:"g1" from b1 show " cktSat nl (A \<leadsto> g1 ->T B)" proof(simp del: cktSat_def,rule_tac fundmental1) show "defSqOfTrForm (True ->T B) \<sqsubseteq>sq defTrajOfCirc A nl" proof(simp add:Let_def) have "(cktSat nl (A\<leadsto> B))" by(cut_tac a3 b1,simp) then show "defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A nl" proof(rule_tac fundmental2) qed qed qed next assume b1:"¬ g1" from b1 show ?thesis by simp qed thm impI lemma steImpI1: assumes a3:"g1 ==> (cktSat nl (A\<leadsto> B))" and a1:"nl:netlists" and a2:"isClosed nl" shows " cktSat nl (A \<leadsto> g1->T B)" proof(case_tac g1) assume b1:"g1" from b1 show " cktSat nl (A \<leadsto> g1 ->T B)" proof(simp del: cktSat_def,rule_tac fundmental1) show "defSqOfTrForm (True ->T B) \<sqsubseteq>sq defTrajOfCirc A nl" proof(simp add:Let_def) have "(cktSat nl (A\<leadsto> B))" by(cut_tac a3 b1,simp) then show "defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A nl" proof(rule_tac fundmental2) qed qed qed next assume b1:"¬ g1" from b1 show ?thesis by simp qed lemma steGuardI: assumes a3:" cktSat nl (A \<leadsto> g1->T B)" and a1:"nl:netlists" and a2:"isClosed nl" shows "g1 --> (cktSat nl (A\<leadsto> B))" proof(rule impI) assume b1:"g1" show "(cktSat nl (A\<leadsto> B))" proof(rule fundmental1) show " defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A nl" proof - from a3 b1 have "cktSat nl (A \<leadsto> True->T B)" by simp then have "defSqOfTrForm ( True->T B) \<sqsubseteq>sq defTrajOfCirc A nl" by(rule_tac fundmental2) then show " defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A nl" proof (simp add:Let_def) qed qed qed qed lemma steGConsStrengthen: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"cktSat nl (A\<leadsto> B)" and a4:"cktSat nl (C\<leadsto> g ->T A)" shows "cktSat nl (C\<leadsto> g ->T B)" proof(rule steImpI) from a4 have "g --> cktSat nl (C\<leadsto> A)" proof(rule steGuardI)qed with a3 show "g --> cktSat nl (C \<leadsto> B)" proof(auto) qed qed lemma steGConsConjI: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"cktSat nl (A\<leadsto> g ->T B )" and a4:"cktSat nl (A\<leadsto> g ->T C )" shows " (cktSat nl (A\<leadsto> g ->T (B andT C)))" proof(rule steImpI) show "g --> cktSat nl (A \<leadsto> B andT C)" proof(rule impI) assume b1:"g" show "cktSat nl (A \<leadsto> B andT C)" proof(rule steConjI,(assumption)+) from a3 have "g -->cktSat nl (A\<leadsto> B )" proof(rule steGuardI) qed with b1 show "cktSat nl (A \<leadsto> B)" by simp next from a4 have "g -->cktSat nl (A\<leadsto> C )" proof(rule steGuardI) qed with b1 show "cktSat nl (A \<leadsto> C)" by simp qed qed qed lemma guardWeakenConjI: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"cktSat nl (A\<leadsto> g2 ->T B )" and a4:"g1 --> g2" shows "cktSat nl (A\<leadsto> g1 ->T B )" proof(rule steImpI) have "g2 --> cktSat nl (A \<leadsto> B)" proof(cut_tac a1,rule steGuardI) qed with a4 show "g1 --> cktSat nl (A \<leadsto> B)" by simp qed lemma steCongAnt: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"cktSat nl (A \<leadsto> B)" and a4:"defSqOfTrForm A' =defSqOfTrForm A" shows "cktSat nl (A' \<leadsto> B)" proof(rule fundmental1) have b0:" defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A nl" proof(rule fundmental2) qed with a4 have "defTrajOfCirc A nl\<sqsubseteq>sq defTrajOfCirc A' nl" thm fSeqIsMono1 proof(unfold defTrajOfCirc_def,rule_tac fSeqIsMono1, auto intro:sqLeqIsReflexive) qed with b0 show c1:" defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A' nl" by (auto intro:sqLeqIsTrans) qed thm mp lemma steCongCons: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"cktSat nl (A \<leadsto> B')" and a4:"defSqOfTrForm B =defSqOfTrForm B'" shows "cktSat nl (A \<leadsto> B)" proof(rule fundmental1) show "defSqOfTrForm B \<sqsubseteq>sq defTrajOfCirc A nl" proof - have b1:"defSqOfTrForm B' \<sqsubseteq>sq defTrajOfCirc A nl" proof(rule fundmental2)qed with a4 show ?thesis by simp qed qed lemma steMp: assumes a1:"g" shows "defSqOfTrForm (g->T B) =defSqOfTrForm B" proof(cut_tac a1,simp add:Let_def) qed lemma steAndComm: shows "defSqOfTrForm (A andT B) =defSqOfTrForm (B andT A)" proof(simp add:Let_def lubComm) qed lemma steAndJiehe: shows "defSqOfTrForm ((A andT B) andT C) =defSqOfTrForm (A andT (B andT C))" proof(simp add:Let_def lubJieHe) qed lemma ElimFalseGuard: shows "defSqOfTrForm (A andT (False ->T B)) =defSqOfTrForm ( A)" proof(simp add:Let_def lub_def) qed lemma ElimTrueGuard: shows "defSqOfTrForm ( (True ->T B)) =defSqOfTrForm ( B)" proof(simp add:Let_def lub_def) qed lemma andCong: assumes a1:"defSqOfTrForm A =defSqOfTrForm A'" shows "defSqOfTrForm (A andT B) =defSqOfTrForm ( A' andT B )" proof(cut_tac a1,simp add:Let_def ) qed lemma andChaosId: shows "defSqOfTrForm (A andT chaos) =defSqOfTrForm A" proof(simp add:Let_def,(rule ext)+,simp add:lub_def X_def Pair_fst_snd_eq) qed constdefs Isb::"node => bool=> trajForm" "Isb n bp== (bp ->T Is1 n) andT ((¬ bp) ->T Is0 n)" (*have proof(rule_tac ?sq1.0=" sqLeqIsTrans:*) lemma steDelAnt: "[|onNodes A ∩ defNodes (InducedNetByNames nl (onNodes C)) = {}; isClosed nl; nl ∈ netlists; ∀n. n ∈ onNodes C --> isDefinedIn n nl;cktSat nl (B \<leadsto> C)|] ==> cktSat nl (B andT A \<leadsto> C) " proof(auto dest:independentNodes) qed constdefs nodeSetSym :: "(node => node) => node set => node set => entity set => bool" "nodeSetSym f M N nl ≡ sym f (InducedNetByNames nl M) (InducedNetByNames nl N)" lemma steSubsetI1: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isClosed nl0" and a5:"isClosed nl" and a6:"ALL n. n:( onNodes C) --> isDefinedIn n nl0" and a7:"(cktSat nl (A \<leadsto> C)) " shows " (cktSat nl0 (A \<leadsto> C))" proof(cut_tac steSubsetI a1 a2 a3 a4 a6 a7,force)qed lemma steSubsetI2: assumes a1:"nl:netlists" and a2:"nl0:netlists" and a3:"nl0 ⊆ nl" and a4:"isClosed nl0" and a7:"isClosed nl" and a5:"ALL n. n:( onNodes C) --> isDefinedIn n nl0" and a6:"(cktSat nl0 (A \<leadsto> C)) " shows "(cktSat nl (A \<leadsto> C))" proof(cut_tac steSubsetI[where nl="nl" ] a1 a2 a3 a4 a5 a6 a7,force)qed lemma steSymReduce2: assumes a1:"isClosed nl" and a2:" nl ∈ netlists" and a3"cktSat nl (A \<leadsto> C)" and a4:"∀ n. n: (onNodes C) --> isDefinedIn n nl" and a5:"∀ n. n: (onNodes (applySym2Form f C)) --> isDefinedIn n nl" and a7:"nodeSetSym f (onNodes C) (onNodes (applySym2Form f C)) nl" and a8:"isSwap f" shows "cktSat nl ((applySym2Form f A) \<leadsto> (applySym2Form f C))" proof - let ?nl="(InducedNetByNames nl (onNodes C))" let ?ml="(InducedNetByNames nl (onNodes (applySym2Form f C)))" have c0a:"{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n} ⊆ nl" apply(rule subsetI) apply simp apply(erule exE) apply(unfold isDefinedIn_def) thm conjE apply(erule conjE)+ apply(erule bexE) apply(cut_tac a2) apply(drule lookUpDevName) apply simp apply auto done have c0b: "{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes (applySym2Form f C) ∧ g = lookUp nl n} ⊆ nl" apply(rule subsetI) apply simp apply(erule exE) apply(unfold isDefinedIn_def) thm conjE apply(erule conjE)+ apply(erule bexE) apply(cut_tac a2) apply(drule lookUpDevName) apply simp apply auto done have b1:"cktSat ?nl (A \<leadsto> C)" proof(rule_tac ?nl="nl" in steSubsetI1) show c1a: "InducedNetByNames nl (onNodes C) ∈ netlists" proof(unfold InducedNetByNames_def, rule subNetIsNetlist, assumption, cut_tac c0a, simp) qed show d1a:" InducedNetByNames nl (onNodes C) ⊆ nl" apply(unfold InducedNetByNames_def) apply(blast intro:subsetNetRef) done show d1b:" isClosed (InducedNetByNames nl (onNodes C))" apply(unfold InducedNetByNames_def) apply(rule subsetNetIsClosed) apply assumption+ apply(cut_tac c0a,simp) done show " ∀n. n ∈ onNodes C --> isDefinedIn n (InducedNetByNames nl (onNodes C))" proof(rule allI,rule impI,unfold InducedNetByNames_def) fix n assume d1:"n ∈ onNodes C" have d2:"{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n } ⊆ (subNetList nl {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n})" proof(rule subsetSubNet2) show " {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n} ⊆ nl" proof fix x assume e1:"x ∈ {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n}" show "x:nl" proof(cut_tac e1,simp) assume f1:" ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ x = lookUp nl n" (is "EX n. ?P n") from f1 obtain na where f2:"?P na" by blast then show "x:nl" proof(unfold isDefinedIn_def) assume g1:"(∃l∈nl. fanout l = na) ∧ na ∈ onNodes C ∧ x = lookUp nl na" from g1 have g2:"(∃l∈nl. fanout l = na)" by simp then obtain l where g3:"l∈nl & ( na=fanout l ) " by auto with g1 have g4:" x = lookUp nl ( fanout l )" by simp have g5:" lookUp nl ( fanout l ) =l " proof(cut_tac g3,rule lookUpDevName)qed(auto) with g4 have "x=l" by simp with g3 show "x ∈ nl" by simp qed qed qed qed have d3:"EX g. g ∈ {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n} & g = lookUp nl n" (is "EX g. ?R g") proof(cut_tac d1 a4,force)qed then obtain g where d4:" ?R g" by blast have d4a:"n=fanout g" thm lookUpFanOut proof(rule sym,cut_tac d4, simp, rule lookUpFanOut) show "isDefinedIn n nl" proof( cut_tac a4 d1, force)qed qed have d5:"g ∈(subNetList nl {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n})" proof(cut_tac d2 d4,blast)qed show "isDefinedIn n (subNetList nl {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes C ∧ g = lookUp nl n})" proof(cut_tac d5 d4a,simp,rule outOfDevIsDefined )qed qed qed have b2:" cktSat (InducedNetByNames nl (onNodes (applySym2Form f C))) ((applySym2Form f A) \<leadsto> (applySym2Form f C) )" proof(rule_tac ?M="?nl" in symReduction1 ) show "sym.sym f (InducedNetByNames nl (onNodes C)) (InducedNetByNames nl (onNodes (applySym2Form f C)))" proof(cut_tac a7, unfold nodeSetSym_def)qed show " InducedNetByNames nl (onNodes (applySym2Form f C)) ∈ netlists" proof(unfold InducedNetByNames_def, rule subNetIsNetlist) show "{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes (applySym2Form f C) ∧ g = lookUp nl n} ⊆ nl" proof(cut_tac c0b, simp)qed qed next show "InducedNetByNames nl (onNodes C) ∈ netlists" proof(unfold InducedNetByNames_def, rule subNetIsNetlist) show "{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ( C) ∧ g = lookUp nl n} ⊆ nl" proof(cut_tac c0a, simp)qed qed next show " isClosed (InducedNetByNames nl (onNodes C))" proof(cut_tac c0a,unfold InducedNetByNames_def,rule subsetNetIsClosed)qed next show " isClosed (InducedNetByNames nl (onNodes (applySym2Form f C)))" proof(cut_tac c0b,unfold InducedNetByNames_def,rule subsetNetIsClosed)qed next show "cktSat (InducedNetByNames nl (onNodes C)) (A \<leadsto> C)" proof(cut_tac b1,simp)qed qed from b2 show ?thesis thm steSubsetI2 proof(rule_tac ?nl="nl" and ?nl0.0="?ml" in steSubsetI2) show " InducedNetByNames nl (onNodes (applySym2Form f C)) ∈ netlists" proof(unfold InducedNetByNames_def, rule subNetIsNetlist) show "{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes (applySym2Form f C) ∧ g = lookUp nl n} ⊆ nl" proof(cut_tac c0b, simp)qed qed next show " isClosed (InducedNetByNames nl (onNodes (applySym2Form f C)))" proof(cut_tac c0b,unfold InducedNetByNames_def,rule subsetNetIsClosed)qed next let ?C="(applySym2Form f C)" show " ∀n. n ∈ onNodes ?C --> isDefinedIn n (InducedNetByNames nl (onNodes ?C))" proof(rule allI,rule impI,unfold InducedNetByNames_def) fix n assume d1:"n ∈ onNodes ?C" have d2:"{g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ g = lookUp nl n } ⊆ (subNetList nl {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ g = lookUp nl n})" proof(rule subsetSubNet2) show " {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ g = lookUp nl n} ⊆ nl" proof fix x assume e1:"x ∈ {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ g = lookUp nl n}" show "x:nl" proof(cut_tac e1,simp) assume f1:" ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ x = lookUp nl n" (is "EX n. ?P n") from f1 obtain na where f2:"?P na" by blast then show "x:nl" proof(unfold isDefinedIn_def) assume g1:"(∃l∈nl. fanout l = na) ∧ na ∈ onNodes ?C ∧ x = lookUp nl na" from g1 have g2:"(∃l∈nl. fanout l = na)" by simp then obtain l where g3:"l∈nl & ( na=fanout l ) " by auto with g1 have g4:" x = lookUp nl ( fanout l )" by simp have g5:" lookUp nl ( fanout l ) =l " proof(cut_tac g3,rule lookUpDevName)qed(auto) with g4 have "x=l" by simp with g3 show "x ∈ nl" by simp qed qed qed qed have d3:"EX g. g ∈ {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ g = lookUp nl n} & g = lookUp nl n" (is "EX g. ?R g") proof(cut_tac d1 a5,auto)qed then obtain g where d4:" ?R g" by blast have d4a:"n=fanout g" thm lookUpFanOut proof(rule sym,cut_tac d4, simp, rule lookUpFanOut) show "isDefinedIn n nl" proof( cut_tac a5 d1, force)qed qed have d5:"g ∈(subNetList nl {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ g = lookUp nl n})" proof(cut_tac d2 d4,blast)qed show "isDefinedIn n (subNetList nl {g. ∃n. isDefinedIn n nl ∧ n ∈ onNodes ?C ∧ g = lookUp nl n})" proof(cut_tac d5 d4a,simp,rule outOfDevIsDefined )qed qed show d1a:" InducedNetByNames nl (onNodes ?C) ⊆ nl" apply(unfold InducedNetByNames_def) apply(blast intro:subsetNetRef) done show " cktSat (InducedNetByNames nl (onNodes (applySym2Form f C))) (applySym2Form f A \<leadsto> applySym2Form f C)" proof(cut_tac b2,simp)qed qed qed end
lemma outOfDevIsDefined:
l ∈ nl ==> isDefinedIn (fanout l) nl
lemma nodeNameIsUniqueInNet2:
[|nl ∈ netlists; isDefinedIn n nl|] ==> ∃!enttr. enttr ∈ nl ∧ fanout enttr = n
lemma isDeifinedInSub:
isDefinedIn m (subNetList nl nl') ==> isDefinedIn m nl
lemma isDeifinedInSub1:
[|isDefinedIn m nl'; nl' ⊆ nl|] ==> isDefinedIn m nl
lemma lookUpSub1:
[|nl' ⊆ nl; isDefinedIn n nl'; nl ∈ netlists; nl' ∈ netlists|] ==> lookUp nl n = lookUp nl' n
lemma lookUpSub:
[|nl' ⊆ nl; isDefinedIn n (subNetList nl nl'); nl ∈ netlists|] ==> lookUp nl n = lookUp (subNetList nl nl') n
lemma lookUpDevName:
[|nl ∈ netlists; enttr ∈ nl|] ==> lookUp nl (fanout enttr) = enttr
lemma lookUpFanOut:
[|nl ∈ netlists; isDefinedIn n nl|] ==> fanout (lookUp nl n) = n
lemma subsetNetIsClosed:
[|nl ∈ netlists; isClosed nl; nl' ⊆ nl|] ==> isClosed (subNetList nl nl')
lemma independentAux1:
∀sq sq0. (∀t n. n ∈ onNodes A --> sq t n = sq0 t n) --> sq \<Turnstile> A = sq0 \<Turnstile> A
lemma independentAux2:
∀t n. n ∉ onNodes A --> defSqOfTrForm A t n = X
lemma independentAux2:
n ∉ onNodes A ==> defSqOfTrForm A t n = X
lemma independentAux3:
[|nl ∈ netlists; ∀t n. isDefinedIn n nl --> sq0.0 t n = sq t n; isClosed nl|] ==> ∀t n. isDefinedIn n nl --> fSeq nl sq0.0 t n = fSeq nl sq t n
lemma XIsLubUnit:
lub a X = a
lemma rclosureSubNet:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; x ∈ rclosure nl0.0 s|] ==> x ∈ rclosure nl s
lemma pairInzipList:
∀stateLs. pair mem zip inps stateLs --> fst pair mem inps
lemma pairInzipList:
pair mem zip inps stateLs ==> fst pair mem inps
lemma GateInNetlist:
[|nl ∈ netlists; m mem inps|] ==> Gate n inps tab ∈ nl --> isDefinedIn m nl
lemma GateInNetlist:
[|Gate n inps tab ∈ nl; nl ∈ netlists; m mem inps|] ==> isDefinedIn m nl
lemma rclosureSubNet':
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl|] ==> x ∈ rclosure nl s --> isDefinedIn (fst x) nl0.0 --> x ∈ rclosure nl0.0 s
lemma rclosureSubNet':
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; x ∈ rclosure nl s; isDefinedIn (fst x) nl0.0|] ==> x ∈ rclosure nl0.0 s
lemma fclosureSubNet:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isDefinedIn n nl0.0|] ==> fclosure nl s n = fclosure nl0.0 s n
lemma curFunEqAny:
[|nl ∈ netlists; ∀n. isDefinedIn n nl --> s1.0 n = s2.0 n; isDefinedIn n nl|] ==> fclosure nl s1.0 n = fclosure nl s2.0 n
lemma fSeqSubNet:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isClosed nl0.0|] ==> ∀n. isDefinedIn n nl0.0 --> fSeq nl sq t n = fSeq nl0.0 sq t n
lemma fSeqSubNet:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isClosed nl0.0; isDefinedIn n nl0.0|] ==> fSeq nl sq t n = fSeq nl0.0 sq t n
lemma trajOfFormInSubNet:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isClosed nl0.0|] ==> ∀n. isDefinedIn n nl0.0 --> defTrajOfCirc B nl t n = defTrajOfCirc B nl0.0 t n
lemma trajOfFormInSubNet:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isClosed nl0.0; isDefinedIn n nl0.0|] ==> defTrajOfCirc B nl t n = defTrajOfCirc B nl0.0 t n
lemma steSubsetI:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isClosed nl0.0; isClosed nl; ∀n. n ∈ onNodes C --> isDefinedIn n nl0.0|] ==> cktSat nl (A \<leadsto> C) = cktSat nl0.0 (A \<leadsto> C)
lemma subsetSubNet2:
[|nl ∈ netlists; nl0.0 ⊆ nl|] ==> nl0.0 ⊆ subNetList nl nl0.0
lemma independentNodes:
[|onNodes A ∩ defNodes (InducedNetByNames nl (onNodes C)) = {}; isClosed nl; nl ∈ netlists; ∀n. n ∈ onNodes C --> isDefinedIn n nl|] ==> cktSat nl (B andT A \<leadsto> C) = cktSat nl (B \<leadsto> C)
lemma GateNodeFclosureValaux1:
nl ∈ netlists ==> (∀m. m mem inps --> isDefinedIn m nl) --> pair0.0 mem zip inps (map (fclosure nl s) inps) --> pair0.0 ∈ rclosure nl s
lemma GateNodeFclosureValaux1:
[|nl ∈ netlists; ∀m. m mem inps --> isDefinedIn m nl; pair0.0 mem zip inps (map (fclosure nl s) inps)|] ==> pair0.0 ∈ rclosure nl s
lemma GateNodeFclosureVal:
[|isClosed nl; Gate out inps tab ∈ nl; nl ∈ netlists; stateLs = map (fclosure nl s) inps; ∀l. l ∈ set tab --> length l = length inps|] ==> fclosure nl s out = lub (funOfTab tab stateLs) (s out)
lemma GateNodeVal:
[|isClosed nl; Gate n inps tab ∈ nl; nl ∈ netlists; ∀l. l ∈ set tab --> length l = length inps|] ==> funOfTab tab (map s inps) \<sqsubseteq> fclosure nl s n
lemma aux1:
T \<sqsubseteq> defSqOfTrForm (andLists (map Is1 (a # list))) 0 a
lemma aux2:
∀n. s' n \<sqsubseteq> s n ==> bv mem map s list --> (∃bv0. bv0 \<sqsubseteq> bv ∧ bv0 mem map s' list)
lemma aux3:
∀n. defSqOfTrForm (andLists (map Is1 list)) 0 n \<sqsubseteq> defSqOfTrForm (andLists (map Is1 (a # list))) 0 n
lemma aux2:
[|∀n. s' n \<sqsubseteq> s n; bv mem map s list|] ==> ∃bv0. bv0 \<sqsubseteq> bv ∧ bv0 mem map s' list
lemma andTabPropT:
[|isAndTab tab; Gate out inps tab ∈ nl; nl ∈ netlists; isClosed nl; ∀l. l ∈ set tab --> length l = length inps|] ==> cktSat nl (andLists (map Is1 inps) \<leadsto> Is1 out)
lemma andTabPropF:
[|isAndTab tab; Gate out inps tab ∈ nl; nl ∈ netlists; isClosed nl; ∀l. l ∈ set tab --> length l = length inps; inpsi mem inps|] ==> cktSat nl (Is0 inpsi \<leadsto> Is0 out)
lemma listLeqstate:
listLeq (map s zs) (map (λu. lub (if u = z then v else X) (s u)) zs)
lemma PosAss2PosVal:
∀inps. length line = length inps --> listLeq (PosValOfLine line) (map (defSqOfTrForm (andLists (PosAssertOfLine inps line)) 0) inps)
lemma PosAss2PosVal:
length line = length inps ==> listLeq (PosValOfLine line) (map (defSqOfTrForm (andLists (PosAssertOfLine inps line)) 0) inps)
lemma orTabPropT:
[|Gate out inps tab ∈ nl; nl ∈ netlists; isClosed nl; ∀l. l ∈ set tab --> length l = length inps; line mem tab|] ==> cktSat nl (andLists (PosAssertOfLine inps line) \<leadsto> Is1 out)
lemma aux1:
∀line. pair mem zip inps line --> (snd pair, f (fst pair)) mem zip line (map f inps)
lemma aux1:
pair mem zip inps line ==> (snd pair, f (fst pair)) mem zip line (map f inps)
lemma aux2:
[|pair mem zip inps line; snd pair = ZERO; fst pair = n; contradict (snd pair, f (fst pair))|] ==> T \<sqsubseteq> f n
lemma aux2:
∀n. s' n \<sqsubseteq> s n ==> ∀line. length line = length inps --> asOfVal mem zip line (map s' inps) ∧ contradict asOfVal --> (∃asOfVal. asOfVal mem zip line (map s inps) ∧ contradict asOfVal)
lemma aux2:
[|∀n. s' n \<sqsubseteq> s n; length line = length inps; asOfVal mem zip line (map s' inps); contradict asOfVal|] ==> ∃asOfVal. asOfVal mem zip line (map s inps) ∧ contradict asOfVal
lemma aux2:
[|asOfVal mem zip line (map (defSqOfTrForm (andLists list) 0) inps) ∧ contradict asOfVal; length line = length inps|] ==> ∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists (aa # list)) 0) inps) ∧ contradict asOfVal
lemma negAux1:
Is1 n mem asLists --> n mem inps --> length line = length inps --> (∃pair. pair mem zip inps line ∧ fst pair = n ∧ snd pair = ZERO) --> (∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)
lemma negAux1:
[|Is1 n mem asLists; n mem inps; length line = length inps; ∃pair. pair mem zip inps line ∧ fst pair = n ∧ snd pair = ZERO|] ==> ∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal
lemma negAux2:
Is0 n mem asLists --> n mem inps --> length line = length inps --> (∃pair. pair mem zip inps line ∧ fst pair = n ∧ snd pair = ONE) --> (∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)
lemma negAux2:
[|Is0 n mem asLists; n mem inps; length line = length inps; ∃pair. pair mem zip inps line ∧ fst pair = n ∧ snd pair = ONE|] ==> ∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal
lemma negAss2NegVal:
∀inps. (∃as. as mem asLists ∧ isNegAssertOfLine as inps line) --> length line = length inps --> (∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal)
lemma negAss2NegVal:
[|∃as. as mem asLists ∧ isNegAssertOfLine as inps line; length line = length inps|] ==> ∃asOfVal. asOfVal mem zip line (map (defSqOfTrForm (andLists asLists) 0) inps) ∧ contradict asOfVal
lemma orTabPropF:
[|Gate out inps tab ∈ nl; nl ∈ netlists; isClosed nl; ∀l. l ∈ set tab --> length l = length inps; ∀line. line mem tab --> (∃as. as mem asLists ∧ isNegAssertOfLine as inps line)|] ==> cktSat nl (andLists asLists \<leadsto> Is0 out)
lemma steRef:
[|nl ∈ netlists; isClosed nl|] ==> cktSat nl (A \<leadsto> A)
lemma steConjI:
[|nl ∈ netlists; isClosed nl; cktSat nl (A \<leadsto> B); cktSat nl (A \<leadsto> C)|] ==> cktSat nl (A \<leadsto> B andT C)
lemma steTrans:
[|nl ∈ netlists; isClosed nl; cktSat nl (A \<leadsto> B); cktSat nl (B \<leadsto> C)|] ==> cktSat nl (A \<leadsto> C)
lemma steImpI:
[|g1.0 --> cktSat nl (A \<leadsto> B); nl ∈ netlists; isClosed nl|] ==> cktSat nl (A \<leadsto> g1.0 ->T B)
lemma steImpI1:
[|g1.0 ==> cktSat nl (A \<leadsto> B); nl ∈ netlists; isClosed nl|] ==> cktSat nl (A \<leadsto> g1.0 ->T B)
lemma steGuardI:
[|cktSat nl (A \<leadsto> g1.0 ->T B); nl ∈ netlists; isClosed nl|] ==> g1.0 --> cktSat nl (A \<leadsto> B)
lemma steGConsStrengthen:
[|nl ∈ netlists; isClosed nl; cktSat nl (A \<leadsto> B); cktSat nl (C \<leadsto> g ->T A)|] ==> cktSat nl (C \<leadsto> g ->T B)
lemma steGConsConjI:
[|nl ∈ netlists; isClosed nl; cktSat nl (A \<leadsto> g ->T B); cktSat nl (A \<leadsto> g ->T C)|] ==> cktSat nl (A \<leadsto> g ->T B andT C)
lemma guardWeakenConjI:
[|nl ∈ netlists; isClosed nl; cktSat nl (A \<leadsto> g2.0 ->T B); g1.0 --> g2.0|] ==> cktSat nl (A \<leadsto> g1.0 ->T B)
lemma steCongAnt:
[|nl ∈ netlists; isClosed nl; cktSat nl (A \<leadsto> B); defSqOfTrForm A' = defSqOfTrForm A|] ==> cktSat nl (A' \<leadsto> B)
lemma steCongCons:
[|nl ∈ netlists; isClosed nl; cktSat nl (A \<leadsto> B'); defSqOfTrForm B = defSqOfTrForm B'|] ==> cktSat nl (A \<leadsto> B)
lemma steMp:
g ==> defSqOfTrForm (g ->T B) = defSqOfTrForm B
lemma steAndComm:
defSqOfTrForm (A andT B) = defSqOfTrForm (B andT A)
lemma steAndJiehe:
defSqOfTrForm (A andT B andT C) = defSqOfTrForm (A andT (B andT C))
lemma ElimFalseGuard:
defSqOfTrForm (A andT (False ->T B)) = defSqOfTrForm A
lemma ElimTrueGuard:
defSqOfTrForm (True ->T B) = defSqOfTrForm B
lemma andCong:
defSqOfTrForm A = defSqOfTrForm A' ==> defSqOfTrForm (A andT B) = defSqOfTrForm (A' andT B)
lemma andChaosId:
defSqOfTrForm (A andT chaos) = defSqOfTrForm A
lemma steDelAnt:
[|onNodes A ∩ defNodes (InducedNetByNames nl (onNodes C)) = {}; isClosed nl; nl ∈ netlists; ∀n. n ∈ onNodes C --> isDefinedIn n nl; cktSat nl (B \<leadsto> C)|] ==> cktSat nl (B andT A \<leadsto> C)
lemma steSubsetI1:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isClosed nl0.0; isClosed nl; ∀n. n ∈ onNodes C --> isDefinedIn n nl0.0; cktSat nl (A \<leadsto> C)|] ==> cktSat nl0.0 (A \<leadsto> C)
lemma steSubsetI2:
[|nl ∈ netlists; nl0.0 ∈ netlists; nl0.0 ⊆ nl; isClosed nl0.0; isClosed nl; ∀n. n ∈ onNodes C --> isDefinedIn n nl0.0; cktSat nl0.0 (A \<leadsto> C)|] ==> cktSat nl (A \<leadsto> C)
lemma steSymReduce2:
[|isClosed nl; nl ∈ netlists; a3.0; cktSat nl (A \<leadsto> C); ∀n. n ∈ onNodes C --> isDefinedIn n nl; ∀n. n ∈ onNodes (applySym2Form f C) --> isDefinedIn n nl; nodeSetSym f (onNodes C) (onNodes (applySym2Form f C)) nl; isSwap f|] ==> cktSat nl (applySym2Form f A \<leadsto> applySym2Form f C)