Up to index of Isabelle/HOL/net
theory caseParaComparatortheory caseParaComparator=law: lemma andMono: assumes a1:"nl:netlists" and a2:" isClosed nl" shows "cktSat nl ((A andT B) \<leadsto> A)" proof (rule fundmental1) have "defSqOfTrForm A \<sqsubseteq>sq defSqOfTrForm (A andT B)" proof(simp,unfold stateSqLeq_def stateLeq_def,(rule allI)+,rule lubSelf) qed also have "…\<sqsubseteq>sq defTrajOfCirc (A andT B) nl" proof(unfold defTrajOfCirc_def, rule fSeqIsExtensive)qed ultimately show " defSqOfTrForm A \<sqsubseteq>sq defTrajOfCirc (A andT B) nl" thm leqIsTrans proof(force dest:sqLeqIsTrans)qed qed lemma andListsAllI1: assumes a1:"nl:netlists" and a2:" isClosed nl" shows "(∀ a. (a:set A )--> cktSat nl (ant \<leadsto> a )) --> cktSat nl (ant \<leadsto>(andLists A))" (is "?P A") proof(induct_tac A,simp) fix b list assume b1:"?P list" show "?P (b # list)" proof(rule impI,simp del:cktSat_def) assume c1:" ∀a. (a = b --> cktSat nl (ant \<leadsto> b)) ∧ (a ∈ set list --> cktSat nl (ant \<leadsto> a)) " show "cktSat nl (ant \<leadsto> b andT andLists list)" proof(rule steConjI) from c1 show "cktSat nl (ant \<leadsto> b)" by auto from c1 have "(∀a. a ∈ set list --> cktSat nl (ant \<leadsto> a))" by auto with b1 show "cktSat nl (ant \<leadsto> ( andLists list))" by auto qed qed qed lemma andListsAllI: assumes a1:"nl:netlists" and a2:" isClosed nl" and a3:"(∀ a. (a:set A )--> cktSat nl (ant \<leadsto> a ))" shows " cktSat nl (ant \<leadsto>(andLists A))" proof(cut_tac a1 a2 a3, blast dest:andListsAllI1) qed lemma andListMem: assumes a1:"nl:netlists" and a2:" isClosed nl" shows "( a mem A) --> cktSat nl ((andLists A) \<leadsto> (andLists [a]))" (is "?P a A") proof(induct_tac A) show "a mem [] --> cktSat nl (andLists [] \<leadsto> andLists [a])" by simp next fix aa list assume b0:"?P a list" show "?P a (aa#list)" proof assume b1:"a mem aa # list" show "cktSat nl (andLists (aa # list) \<leadsto> andLists [a])" proof(simp del:cktSat_def) show " cktSat nl (aa andT andLists list \<leadsto>a andT chaos)" proof - from b1 have "a=aa |a mem list" by (simp add:mem_iff) moreover {assume c1:"a=aa" have c2:" cktSat nl (aa andT andLists list \<leadsto> a)" proof(cut_tac c1,simp del:cktSat_def, rule andMono,auto) qed have c3:" cktSat nl (a \<leadsto> andLists [a])" proof(rule andListsAllI) show " ∀aa. aa ∈ set [a] --> cktSat nl (a \<leadsto> aa)" proof(simp del:cktSat_def,auto)qed qed with c2 have " cktSat nl (aa andT andLists list \<leadsto>andLists [a])" proof(rule_tac ?B="a" in steTrans)qed then have " cktSat nl (aa andT andLists list \<leadsto>a andT chaos)" by simp } moreover {assume c1:"a mem list" have c2:" cktSat nl (aa andT andLists list \<leadsto> andLists list)" proof(rule_tac ?B="(andLists list)andT aa" in steTrans) show d1:"cktSat nl (aa andT andLists list \<leadsto>(andLists list)andT aa)" proof(rule steCongCons) show "defSqOfTrForm ((andLists list)andT aa) = defSqOfTrForm (aa andT andLists list)" proof(simp del:cktSat_def only:steAndComm)qed show "cktSat nl (aa andT andLists list \<leadsto> aa andT andLists list)" proof(rule steRef)qed qed show " cktSat nl (andLists list andT aa \<leadsto> andLists list)" proof(rule andMono)qed qed have "cktSat nl (andLists list\<leadsto> a andT chaos)" proof(cut_tac c1 b0,simp del:cktSat_def)qed with c2 have " cktSat nl (aa andT andLists list \<leadsto>a andT chaos)" proof(rule_tac ?B="andLists list" in steTrans) qed } ultimately show " cktSat nl (aa andT andLists list \<leadsto>a andT chaos)" by blast qed qed qed qed lemma andListMem': assumes a1:"nl:netlists" and a2:" isClosed nl" and a3:"( a mem A) " shows " cktSat nl ((andLists A) \<leadsto> (andLists [a]))" proof(cut_tac andListMem a1 a2 a3,auto)qed lemma andListMem1: assumes a1:"nl:netlists" and a2:" isClosed nl" and a3:"( a mem A) " shows " cktSat nl ((andLists A) \<leadsto> (a))" (is "?P a A") proof - have b1:"cktSat nl ((andLists [a]) \<leadsto>a)" proof(simp)qed have "cktSat nl ((andLists A) \<leadsto>(andLists [a]))" proof(rule andListMem') qed with b1 show " cktSat nl ((andLists A) \<leadsto> (a))"thm steTrans proof(rule_tac steTrans)qed qed lemma andListSubsetEq: assumes a1:"nl:netlists" and a2:"isClosed nl" and a3:"((set B) Int - {chaos})⊆ set A" shows "cktSat nl ((andLists A) \<leadsto> (andLists B))" proof(rule andListsAllI) show "∀a. a ∈ set B --> cktSat nl (andLists A \<leadsto> a)" proof(rule allI,rule impI) fix a assume b1:"a ∈ set B" show "cktSat nl (andLists A \<leadsto> a)" proof(case_tac "a=chaos") assume c1:"a = chaos" show "cktSat nl (andLists A \<leadsto> a)" proof(cut_tac c1,rule_tac fundmental1) show "defSqOfTrForm a \<sqsubseteq>sq defTrajOfCirc (andLists A) nl" proof(cut_tac c1,unfold stateSqLeq_def stateLeq_def, (rule allI)+,force intro:XIsLeastValue) qed qed next assume c1:"a ~= chaos" have c2:"a ∈ set A" proof(cut_tac a3 c1 b1,auto) qed then show "cktSat nl (andLists A \<leadsto> a)" proof(rule_tac andListMem1) show " a mem A" proof(cut_tac c2, simp add:mem_iff)qed qed qed qed qed lemma andListCat: assumes a1:"nl:netlists" and a2:" isClosed nl" and a3:"cktSat nl ((andLists A1) \<leadsto> (andLists B1))" and a4:"cktSat nl ((andLists A2) \<leadsto> (andLists B2))" shows "cktSat nl ((andLists (A1@A2)) \<leadsto> (andLists (B1@B2)))" proof(rule andListsAllI) show " ∀a. a ∈ set (B1 @ B2) --> cktSat nl (andLists (A1@ A2) \<leadsto> a)" proof(rule allI, rule impI) fix a assume b1:"a ∈ set (B1 @ B2)" from b1 have b2:"a ∈ set(B1) |a ∈ set B2" by simp moreover {assume b3:"a ∈ set(B1)" have b4:"cktSat nl ((andLists B1) \<leadsto> a)" proof(rule andListMem1) show "a mem B1 " proof(cut_tac b3,simp add:mem_iff)qed qed have b5:"cktSat nl ((andLists A1) \<leadsto> a)" proof(cut_tac b4 a3, force intro:steTrans)qed have "cktSat nl ((andLists(A1 @ A2) ) \<leadsto>andLists(A1))" proof(rule andListsAllI) show " ∀a. a ∈ set A1 --> cktSat nl (andLists (A1 @ A2) \<leadsto> a)" proof(rule allI,rule impI) fix a1 assume c1:"a1 ∈ set A1" show "cktSat nl (andLists (A1 @ A2) \<leadsto> a1)" proof(rule andListMem1) show "a1 mem A1 @ A2" proof(cut_tac c1,simp add:mem_iff) qed qed qed qed with b5 have "cktSat nl ((andLists(A1 @ A2)) \<leadsto>a)" by(force intro:steTrans) } moreover {assume b3:"a ∈ set(B2)" have b4:"cktSat nl ((andLists B2) \<leadsto> a)" proof(rule andListMem1) show "a mem B2 " proof(cut_tac b3,simp add:mem_iff)qed qed have b5:"cktSat nl ((andLists A2) \<leadsto> a)" proof(cut_tac b4 a4, force intro:steTrans)qed have "cktSat nl ((andLists(A1 @ A2)) \<leadsto>andLists(A2))" proof(rule andListsAllI) show " ∀a. a ∈ set A2 --> cktSat nl (andLists (A1 @ A2) \<leadsto> a)" proof(rule allI,rule impI) fix a1 assume c1:"a1 ∈ set A2" show "cktSat nl (andLists (A1 @ A2) \<leadsto> a1)" proof(rule andListMem1) show "a1 mem A1 @ A2" proof(cut_tac c1,simp add:mem_iff) qed qed qed qed with b5 have "cktSat nl ((andLists(A1 @ A2)) \<leadsto>a)" by(force intro:steTrans) } ultimately show "cktSat nl ((andLists(A1 @ A2)) \<leadsto>a)" by auto qed qed thm ballI lemma caseAux1: assumes a1:"xorTab=[[ONE,ONE],[ZERO,ZERO]]" and a2:"andTab=[andLine]" and (* a, b, c: three functions map int to nodes a3:"InputOfAs= {l. EX i<n. l=Input (a i)}" and a4:"InputOfBs= {l. EX i<n. l=Input (b i)}" and*) a5:"xorGLs={ l. EX i<n. l= Gate (c i) [(a i),(b i)] xorTab }" and a6:"andLine=map (λ n. ONE) [0..<n]" and a7:"andG=Gate out cs andTab" and a8:" (xorGLs Un {andG} )⊆ nl " and a9:"nl:netlists" and (*bvOfAs, bvOfBs *) a10:"antOfAs=map (λ i. ( Isb (a i) (bvOfAs i))) [0..<n] " and a11:"antOfBs=map (λ i. ( Isb (b i) (bvOfBs i))) [0..<n]" and a12:"i< n &( bvOfAs i) =( bvOfBs i)" and a14:"cons0=Gp0->T Is0 out " and a15:"cons1=Gp1->T Is1 out " and a16:"cons=cons0 andT cons1" and a17:"isClosed nl" and a18:"0<n" and a19:"cs=map c [0..<n]" and a20:"ant=andLists (antOfAs@antOfBs)" shows "cktSat nl (ant \<leadsto> ( Is1 (c i)))" proof - let ?ant0="(Isb (a i) (bvOfAs i))" let ?ant1="(Isb (b i) (bvOfBs i))" have "bvOfAs i = bvOfBs i" proof(cut_tac a12 ,simp)qed then have "(bvOfAs i) &(bvOfBs i )|~(bvOfAs i)&~(bvOfBs i)" by simp moreover {assume c1:"(bvOfAs i) &(bvOfBs i )" from c1 have d1: "defSqOfTrForm (Isb (a i) (bvOfAs i)) =defSqOfTrForm (Is1 (a i))" proof(erule_tac conjE, unfold Isb_def, simp add: steMp ElimTrueGuard ElimFalseGuard del :defSqAnd_def defSqIs1_def defSqGuard_def) qed from c1 have d2:"defSqOfTrForm (Isb (b i) (bvOfBs i)) =defSqOfTrForm (Is1 (b i))" proof(erule_tac conjE, unfold Isb_def, simp add: steMp ElimTrueGuard ElimFalseGuard del :defSqAnd_def defSqIs1_def defSqGuard_def) qed from d1 d2 have d3:"defSqOfTrForm (?ant0 andT ?ant1)= defSqOfTrForm ((Is1 (a i)) andT (Is1 (b i)))" proof(simp add:andCong steAndComm) qed have d4:"cktSat nl (((Is1 (a i)) andT (Is1 (b i)))\<leadsto> Is1 (c i)) " thm steCongAnt proof(rule_tac ?A="andLists (PosAssertOfLine ?inps ?line)" in steCongAnt) thm orTabPropT have e1:" Gate (c i) [(a i),(b i)] xorTab:nl" proof(cut_tac a12 a5 a8,simp,blast)qed let ?line="[ONE,ONE]" let ?tab=" xorTab" let ?inps="[a i,b i]" show e2:"cktSat nl (andLists (PosAssertOfLine ?inps ?line) \<leadsto> Is1 (c i))" proof(cut_tac a1 a6 e1, rule_tac ?tab="?tab" in orTabPropT,(assumption)+,auto ) qed show "defSqOfTrForm (Is1 (a i) andT Is1 (b i)) = defSqOfTrForm (andLists (PosAssertOfLine [a i, b i] [ONE, ONE]))" proof(rule sym,simp add:Let_def andChaosId del:defSqAnd_def) have "defSqOfTrForm (Is1 (a i) andT (Is1 (b i) andT chaos))= defSqOfTrForm (Is1 (a i) andT Is1 (b i) andT chaos)" by(rule sym,simp del:defSqAnd_def add:steAndJiehe) also have "…=defSqOfTrForm (Is1 (a i) andT Is1 (b i))" by(rule andChaosId) ultimately show " defSqOfTrForm (Is1 (a i) andT (Is1 (b i) andT chaos)) = defSqOfTrForm (Is1 (a i) andT Is1 (b i))" by simp qed qed with d3 have "cktSat nl ((?ant0 andT ?ant1)\<leadsto> Is1 (c i))" proof(rule_tac ?A=" ((Is1 (a i)) andT (Is1 (b i)))" in steCongAnt)qed } moreover {assume c1:"(~bvOfAs i) &(~bvOfBs i )" let ?ant0="(Isb (a i) (bvOfAs i))" let ?ant1="(Isb (b i) (bvOfBs i))" from c1 have d1: "defSqOfTrForm ?ant0=defSqOfTrForm (Is0 (a i))" proof(erule_tac conjE,cut_tac a8, unfold Isb_def,simp add: steMp steAndComm ElimTrueGuard ElimFalseGuard del :defSqAnd_def defSqIs0_def defSqGuard_def) have "defSqOfTrForm ((False ->T Is1 (a i) )andT (True ->T Is0 (a i)))= defSqOfTrForm ((True ->T Is0 (a i) )andT ( False ->T Is1 (a i)))" proof(simp add:steAndComm del :defSqAnd_def defSqIs0_def defSqGuard_def ) qed also have "…= defSqOfTrForm (True ->TIs0 (a i))" proof(simp add: ElimFalseGuard del :defSqAnd_def defSqIs0_def defSqGuard_def) qed also have "…=defSqOfTrForm (Is0 (a i))" proof(simp add: ElimTrueGuard del :defSqAnd_def defSqIs0_def defSqGuard_def) qed finally show "defSqOfTrForm ((False ->T Is1 (a i) )andT (True ->T Is0 (a i)))= defSqOfTrForm (Is0 (a i))" by blast qed from c1 have d2: "defSqOfTrForm ?ant1=defSqOfTrForm (Is0 (b i))" proof(erule_tac conjE,cut_tac a11, unfold Isb_def,simp add: steMp steAndComm ElimTrueGuard ElimFalseGuard del :defSqAnd_def defSqIs0_def defSqGuard_def) have "defSqOfTrForm ((False ->T Is1 (b i) )andT (True ->T Is0 (b i)))= defSqOfTrForm ((True ->T Is0 (b i) )andT ( False ->T Is1 (b i)))" proof(simp add:steAndComm del :defSqAnd_def defSqIs0_def defSqGuard_def ) qed also have "…= defSqOfTrForm (True ->TIs0 (b i))" proof(simp add: ElimFalseGuard del :defSqAnd_def defSqIs0_def defSqGuard_def) qed also have "…=defSqOfTrForm (Is0 (b i))" proof(simp add: ElimTrueGuard del :defSqAnd_def defSqIs0_def defSqGuard_def) qed finally show "defSqOfTrForm ((False ->T Is1 (b i) )andT (True ->T Is0 (b i)))= defSqOfTrForm (Is0 (b i))" by blast qed with d1 have d3:"defSqOfTrForm (?ant0 andT ?ant1)= defSqOfTrForm ((Is0 (a i)) andT (Is0 (b i)))" proof(simp add:andCong steAndComm) qed let ?line="[ZERO,ZERO]" let ?tab=" xorTab" let ?inps="[(a i),(b i)]" have "cktSat nl (((Is0 (a i)) andT (Is0 (b i)))\<leadsto> Is1 (c i))" thm steCongAnt proof(rule_tac ?A="andLists (PosAssertOfLine ?inps ?line)" in steCongAnt) have e1:" Gate (c i) [(a i),(b i)] xorTab:nl" proof(cut_tac a12 a5 a8,simp,blast)qed show e2:"cktSat nl (andLists (PosAssertOfLine ?inps ?line) \<leadsto> Is1 (c i))" proof(cut_tac a1 a6 e1, rule_tac ?tab="?tab" in orTabPropT,(assumption)+,auto ) qed show "defSqOfTrForm (Is0 (a i) andT Is0 (b i)) = defSqOfTrForm (andLists (PosAssertOfLine [a i, b i] ?line))" proof(rule sym,simp add:Let_def andChaosId del:defSqAnd_def) have "defSqOfTrForm (Is0 (a i) andT (Is0 (b i) andT chaos))= defSqOfTrForm (Is0 (a i) andT Is0 (b i) andT chaos)" by(rule sym,simp del:defSqAnd_def add:steAndJiehe) also have "…=defSqOfTrForm (Is0 (a i) andT Is0 (b i))" by(rule andChaosId) ultimately show " defSqOfTrForm (Is0 (a i) andT (Is0 (b i) andT chaos)) = defSqOfTrForm (Is0 (a i) andT Is0 (b i))" by simp qed qed with d3 have "cktSat nl ((?ant0 andT ?ant1)\<leadsto> Is1 (c i))" proof(rule_tac ?A=" ((Is0 (a i)) andT (Is0 (b i)))" in steCongAnt)qed } ultimately have b1:" cktSat nl ((?ant0 andT ?ant1)\<leadsto> Is1 (c i))" by blast have b2:"cktSat nl ((andLists[?ant0 ,?ant1])\<leadsto> (?ant0 andT ?ant1))" proof(simp)qed have b3:" cktSat nl (ant \<leadsto> (andLists [?ant0 ,?ant1]))" proof(cut_tac a10 a11 a20,simp del:cktSat_def andListsCons, rule andListSubsetEq) show " set [Isb (a i) (bvOfAs i), Isb (b i) (bvOfBs i)] ∩ - {chaos} ⊆ set (map (λi. Isb (a i) (bvOfAs i)) [0..<n] @ map (λi. Isb (b i) (bvOfBs i)) [0..<n])" proof(cut_tac a12, simp add:set_upt,force)qed qed with b1 b2 a9 a17 show "cktSat nl (ant \<leadsto>Is1 (c i))" proof(blast dest:steTrans)qed qed lemma caseAux1: assumes a1:"xorTab=[[ONE,ONE],[ZERO,ZERO]]" and a2:"andTab=[andLine]" and (* a, b, c: three functions map int to nodes a3:"InputOfAs= {l. EX i<n. l=Input (a i)}" and a4:"InputOfBs= {l. EX i<n. l=Input (b i)}" and*) a5:"xorGLs={ l. EX i<n. l= Gate (c i) [(a i),(b i)] xorTab }" and a6:"andLine=map (λ n. ONE) [0..<n]" and a7:"andG=Gate out cs andTab" and a8:" (xorGLs Un {andG} )⊆ nl " and a9:"nl:netlists" and (*bvOfAs, bvOfBs *) a10:"antOfAs=map (λ i. ( Isb (a i) (bvOfAs i))) [0..<n] " and a11:"antOfBs=map (λ i. ( Isb (b i) (bvOfBs i))) [0..<n]" and a14:"cons0=Gp0->T Is0 out " and a15:"cons1=Gp1->T Is1 out " and a16:"cons=cons0 andT cons1" and a17:"isClosed nl" and a18:"0<n" and a19:"cs=map c [0..<n]" and a20:"ant=andLists (antOfAs@antOfBs)" and a21:"(ALL i. i< n --> ( bvOfAs i) =( bvOfBs i))" shows "cktSat nl (ant \<leadsto> andLists (map Is1 cs ))" proof(rule andListsAllI) show "∀a. a ∈ set (map Is1 cs) --> cktSat nl (ant \<leadsto> a)" proof(cut_tac a19,simp del:cktSat_def,rule allI,rule impI) fix trf assume b1:" trf∈ Is1 ` c ` {k. k < n}" show "cktSat nl (ant \<leadsto> trf)" proof - from b1 have "EX i. i < n & trf=Is1 (c i)" (is "EX i. ?P i") by blast then obtain i where b2:"?P i" by blast have "cktSat nl (ant \<leadsto> Is1 (c i))" thm caseAux1 proof(rule_tac xorTab="xorTab" and xorGLs="xorGLs" and cs="cs" and ?cons0.0="cons0" and ?cons1.0 ="cons1" and a="a" and b="b" and c="c" and n="n" in caseAux1, (assumption)+) show "i < n ∧ bvOfAs i = bvOfBs i" proof(cut_tac b2 a21,blast)qed qed with b2 show "cktSat nl (ant \<leadsto> trf)" by simp qed qed qed lemma caseComparator: assumes a1:"xorTab=[[ONE,ONE],[ZERO,ZERO]]" and a2:"andTab=[andLine]" and (* a, b, c: three functions map int to nodes a3:"InputOfAs= {l. EX i<n. l=Input (a i)}" and a4:"InputOfBs= {l. EX i<n. l=Input (b i)}" and*) a5:"xorGLs={ l. EX i<n. l= Gate (c i) [(a i),(b i)] xorTab }" and a6:"andLine=map (λ n. ONE) [0..<n]" and a7:"andG=Gate out cs andTab" and a8:" (xorGLs Un {andG} )⊆ nl " and a9:"nl:netlists" and (*bvOfAs, bvOfBs *) a10:"antOfAs=map (λ i. ( Isb (a i) (bvOfAs i))) [0..<n] " and a11:"antOfBs=map (λ i. ( Isb (b i) (bvOfBs i))) [0..<n]" and a12:"Gp0=(EX i. i< n & ( bvOfAs i) ~=( bvOfBs i))" and a13:"Gp1=(ALL i. i< n --> ( bvOfAs i) =( bvOfBs i))" and a14:"cons0=Gp0->T Is0 out " and a15:"cons1=Gp1->T Is1 out " and a16:"cons=cons0 andT cons1" and a17:"isClosed nl" and a18:"0<n" and a19:"cs=map c [0..<n]" and a20:"ant=andLists (antOfAs@antOfBs)" shows "cktSat nl (ant \<leadsto> cons)" proof(cut_tac a16,simp del:cktSat_def,rule steConjI) show "cktSat nl (ant \<leadsto> cons0)" proof(cut_tac a12 a14,simp del:cktSat_def,rule steImpI1) assume b1:" ∃i<n. bvOfAs i = (¬ bvOfBs i)" from b1 obtain i where b2:"i<n & bvOfAs i = (¬ bvOfBs i)" by blast let ?PM="(Is0 (c i))" show " cktSat nl (ant \<leadsto> Is0 out)" proof(rule_tac ?B="?PM" in steTrans) have "(bvOfAs i)&~(bvOfBs i) | ~(bvOfAs i)&(bvOfBs i)" proof(cut_tac b2,blast)qed moreover {assume c1:"(bvOfAs i)&~(bvOfBs i)" let ?line="[ONE,ONE]" let ?tab=" xorTab" let ?inps="[(a i),(b i)]" let ?asLists="[Is1 (a i), Is0 (b i)]" let ?PN="(andLists ?asLists)" have "cktSat nl (ant \<leadsto> Is0 (c i))" proof(rule_tac ?B="?PN" in steTrans) show " cktSat nl (andLists [Is1 (a i), Is0 (b i)] \<leadsto> Is0 (c i))" proof(rule_tac ?inps="?inps" and ?tab="?tab" and ?asLists="?asLists" in orTabPropF) show " ∀line. line mem xorTab --> (∃as. as mem [Is1 (a i), Is0 (b i)] ∧ isNegAssertOfLine as [a i, b i] line)" proof(rule allI, rule impI) fix line assume f1:"line mem xorTab" show " ∃ as. as mem [Is1 (a i), Is0 (b i)] ∧ isNegAssertOfLine as [a i, b i] line" (is "?S") proof - from a1 f1 have "line = [ONE,ONE]| line=[ZERO,ZERO]" proof(simp, case_tac "[ONE, ONE] = line",simp, case_tac "[ZERO, ZERO] = line",auto) qed moreover {assume g1:"line =[ONE,ONE]" have "?S" proof(cut_tac g1,force) qed } moreover {assume g1:"line =[ZERO,ZERO]" have "?S" proof(cut_tac g1,force) qed } ultimately show "?S" by blast qed qed show "∀ l. l ∈ set xorTab --> length l = length [a i, b i]" proof(cut_tac a1,simp) qed show "Gate (c i) [a i, b i] xorTab ∈ nl" proof(cut_tac a5 a8 b2,force)qed qed show " cktSat nl (ant \<leadsto> andLists [Is1 (a i), Is0 (b i)])" proof(rule andListsAllI) show "∀aa. aa ∈ set [Is1 (a i), Is0 (b i)] --> cktSat nl (ant \<leadsto> aa)" proof(rule allI,rule impI) fix aa assume g1:"aa ∈ set [Is1 (a i), Is0 (b i)] " show " cktSat nl (ant \<leadsto> aa)" proof - from g1 have "aa=Is1 (a i) | aa=Is0 (b i)" by simp moreover {assume h1:"aa=Is1 (a i)" have h2:" cktSat nl (ant \<leadsto> Isb ( a i) (bvOfAs i))" proof(cut_tac a10 a11 a20 h1 a18, simp del:cktSat_def, rule andListMem1) have "i<n" by(cut_tac b2,simp) then show "Isb (a i) (bvOfAs i) mem map (λi. Isb (a i) (bvOfAs i)) [0..<n] @ map (λi. Isb (b i) (bvOfBs i)) [0..<n]" proof(simp add:set_map set_upt mem_iff)qed qed have " cktSat nl (( Isb ( a i)(bvOfAs i))\<leadsto>aa)" proof(cut_tac h1 c1, unfold Isb_def,simp)qed with h2 have " cktSat nl (ant \<leadsto> aa)"thm steTrans proof(rule_tac steTrans) qed } moreover {assume h1:"aa=Is0 (b i)" have h2:" cktSat nl (ant \<leadsto> Isb ( b i) (bvOfBs i))" proof(cut_tac a10 a11 a20 h1 a18, simp del:cktSat_def, rule andListMem1) have "i<n" by(cut_tac b2,simp) then show "Isb (b i) (bvOfBs i) mem map (λi. Isb (a i) (bvOfAs i)) [0..<n] @ map (λi. Isb (b i) (bvOfBs i)) [0..<n]" proof(simp add:set_map set_upt mem_iff)qed qed have " cktSat nl (( Isb ( b i)(bvOfBs i))\<leadsto>aa)" proof(cut_tac h1 c1, unfold Isb_def,simp)qed with h2 have " cktSat nl (ant \<leadsto> aa)"thm steTrans proof(rule_tac steTrans) qed } ultimately show "cktSat nl (ant \<leadsto> aa)" by blast qed qed qed qed } moreover {assume c1:"~(bvOfAs i)&(bvOfBs i)" let ?line="[ONE,ONE]" let ?tab=" xorTab" let ?inps="[(a i),(b i)]" let ?asLists="[Is0 (a i), Is1 (b i)]" let ?PN="(andLists ?asLists)" have "cktSat nl (ant \<leadsto> Is0 (c i))" proof(rule_tac ?B="?PN" in steTrans) show " cktSat nl (andLists [Is0 (a i), Is1 (b i)] \<leadsto> Is0 (c i))" proof(rule_tac ?inps="?inps" and ?tab="?tab" and ?asLists="?asLists" in orTabPropF) show " ∀line. line mem xorTab --> (∃as. as mem [Is0 (a i), Is1 (b i)] ∧ isNegAssertOfLine as [a i, b i] line)" proof(rule allI, rule impI) fix line assume f1:"line mem xorTab" show " ∃ as. as mem [Is0 (a i), Is1 (b i)] ∧ isNegAssertOfLine as [a i, b i] line" (is "?S") proof - from a1 f1 have "line = [ONE,ONE]| line=[ZERO,ZERO]" proof(simp, case_tac "[ONE, ONE] = line",simp, case_tac "[ZERO, ZERO] = line",auto) qed moreover {assume g1:"line =[ONE,ONE]" have "?S" proof(cut_tac g1,force) qed } moreover {assume g1:"line =[ZERO,ZERO]" have "?S" proof(cut_tac g1,force) qed } ultimately show "?S" by blast qed qed show "∀ l. l ∈ set xorTab --> length l = length [a i, b i]" proof(cut_tac a1,simp) qed show "Gate (c i) [a i, b i] xorTab ∈ nl" proof(cut_tac a5 a8 b2,force)qed qed show " cktSat nl (ant \<leadsto> andLists [Is0 (a i), Is1 (b i)])" proof(rule andListsAllI) show "∀aa. aa ∈ set [Is0 (a i), Is1 (b i)] --> cktSat nl (ant \<leadsto> aa)" proof(rule allI,rule impI) fix aa assume g1:"aa ∈ set [Is0 (a i), Is1 (b i)] " show " cktSat nl (ant \<leadsto> aa)" proof - from g1 have "aa=Is0 (a i) | aa=Is1 (b i)" by simp moreover {assume h1:"aa=Is0 (a i)" have h2:" cktSat nl (ant \<leadsto> Isb ( a i) (bvOfAs i))" proof(cut_tac a10 a11 a20 h1 a18, simp del:cktSat_def, rule andListMem1) have "i<n" by(cut_tac b2,simp) then show "Isb (a i) (bvOfAs i) mem map (λi. Isb (a i) (bvOfAs i)) [0..<n] @ map (λi. Isb (b i) (bvOfBs i)) [0..<n]" proof(simp add:set_map set_upt mem_iff)qed qed have " cktSat nl (( Isb ( a i)(bvOfAs i))\<leadsto>aa)" proof(cut_tac h1 c1, unfold Isb_def,simp)qed with h2 have " cktSat nl (ant \<leadsto> aa)"thm steTrans proof(rule_tac steTrans) qed } moreover {assume h1:"aa=Is1 (b i)" have h2:" cktSat nl (ant \<leadsto> Isb ( b i) (bvOfBs i))" proof(cut_tac a10 a11 a20 h1 a18, simp del:cktSat_def, rule andListMem1) have "i<n" by(cut_tac b2,simp) then show "Isb (b i) (bvOfBs i) mem map (λi. Isb (a i) (bvOfAs i)) [0..<n] @ map (λi. Isb (b i) (bvOfBs i)) [0..<n]" proof(simp add:set_map set_upt mem_iff)qed qed have " cktSat nl (( Isb ( b i)(bvOfBs i))\<leadsto>aa)" proof(cut_tac h1 c1, unfold Isb_def,simp)qed with h2 have " cktSat nl (ant \<leadsto> aa)"thm steTrans proof(rule_tac steTrans) qed } ultimately show "cktSat nl (ant \<leadsto> aa)" by blast qed qed qed qed } ultimately show "cktSat nl (ant \<leadsto> Is0 (c i))" by blast let ?inps="cs" let ?andTab="andTab" let ?out="out" show " cktSat nl (Is0 (c i) \<leadsto> Is0 out)" proof(cut_tac a2,rule_tac ?tab="andTab" and inps="?inps"in andTabPropF) show "isAndTab andTab" proof(cut_tac a2 a6 ,unfold isAndTab_def isFullAndLine_def, simp,induct_tac n,auto) fix n l assume b1:" ∀l. l mem map (λn. ONE) [0..<n] --> l = ONE" and b2:" l mem map (λn. ONE) [0..<n] @ [ONE]" have "map (λn. ONE) [0..<n] =[] |map (λn. ONE) [0..<n] ~=[] " by auto moreover {assume c1:"map (λn. ONE) [0..<n] =[]" from b2 c1 have c1: "l=ONE" apply - apply simp apply(case_tac " ONE = l",auto) done } moreover {assume c1:"map (λn. ONE) [0..<n] ~=[]" from b2 have c2:"l : set ( map (λn. ONE) [0..<n]@ [ONE])" by (simp add:mem_iff) from c2 have c2: " l : set (map (λn. ONE) [0..<n]) | l=ONE" apply - apply auto done then have "l=ONE" apply - apply(rule disjE) apply simp apply(cut_tac b1,blast) by simp } ultimately show " l = ONE" by blast qed show "∀ l. l ∈ set andTab --> length l = length ?inps" proof(cut_tac a2 a6 a19,simp) qed show "Gate out cs andTab ∈ nl" proof(cut_tac a7 a8,simp) qed show " c i mem cs" proof(cut_tac a2 b2 a6 a18 a19,simp add:mem_iff) qed qed(cut_tac a5 a6 a7 ,auto) qed qed show "cktSat nl (ant \<leadsto> cons1)" proof(cut_tac a15,simp del:cktSat_def,rule steImpI1) assume b0:"Gp1" let ?line="andLine" let ?tab=" andTab" let ?inps="cs" let ?PM="andLists (map Is1 cs) " (*andLists (PosAssertOfLine ?inps ?line) ((andLists (map Is1 cs)))"*) show "cktSat nl (ant \<leadsto> Is1 out)" proof(rule_tac ?B="?PM" in steTrans) show "cktSat nl (ant \<leadsto> ?PM)" proof(rule_tac cs="cs" in caseAux1,(assumption)+, cut_tac b0 a13,simp)qed let ?inps="cs" let ?andTab="andTab" let ?out="out" show " cktSat nl (andLists (map Is1 cs) \<leadsto> Is1 out)" proof(rule_tac ?tab="andTab" and inps="?inps" in andTabPropT) show "isAndTab andTab" proof(cut_tac a2 a6 ,unfold isAndTab_def isFullAndLine_def, simp,induct_tac n,auto) fix n l assume b1:" ∀l. l mem map (λn. ONE) [0..<n] --> l = ONE" and b2:" l mem map (λn. ONE) [0..<n] @ [ONE]" have "map (λn. ONE) [0..<n] =[] |map (λn. ONE) [0..<n] ~=[] " by auto moreover {assume c1:"map (λn. ONE) [0..<n] =[]" from b2 c1 have c1: "l=ONE" apply - apply simp apply(case_tac " ONE = l",auto) done } moreover {assume c1:"map (λn. ONE) [0..<n] ~=[]" from b2 have c2:"l : set ( map (λn. ONE) [0..<n]@ [ONE])" by (simp add:mem_iff) from c2 have c2: " l : set (map (λn. ONE) [0..<n]) | l=ONE" apply - apply auto done then have "l=ONE" apply - apply(rule disjE) apply simp apply(cut_tac b1,blast) by simp } ultimately show " l = ONE" by blast qed show "∀ l. l ∈ set andTab --> length l = length ?inps" proof(cut_tac a2 a6 a19,simp) qed show "Gate out cs andTab ∈ nl" proof(cut_tac a7 a8,simp) qed qed qed qed qed lemma caseParaComparatorNeq: assumes a1:"xorTab=[[ONE,ONE],[ZERO,ZERO]]" and a2:"andTab=[andLine]" and (* a, b, c: three functions map int to nodes a3:"InputOfAs= {l. EX i<n. l=Input (a i)}" and a4:"InputOfBs= {l. EX i<n. l=Input (b i)}" and*) a5:"xorGLs={ l. EX i<n. l= Gate (c i) [(a i),(b i)] xorTab }" and a6:"andLine=map (λ n. ONE) [0..<n]" and a7:"andG=Gate out cs andTab" and a8:" (xorGLs Un {andG} )⊆ nl " and a9:"nl:netlists" and (*bvOfAs, bvOfBs *) a10:"antOfAs=map (λ i. ( Isb (a i) (bvOfAs i))) [0..<n] " and a11:"antOfBs=map (λ i. ( Isb (b i) (bvOfBs i))) [0..<n]" and a12:"Gp0=(EX i. i< n & ( bvOfAs i) ~=( bvOfBs i))" and a13:"Gp1=(ALL i. i< n --> ( bvOfAs i) =( bvOfBs i))" and a14:"cons0=Gp0->T Is0 out " and a15:"cons1=Gp1->T Is1 out " and a16:"cons=cons0 andT cons1" and a17:"isClosed nl" and a18:"0<n" and a19:"cs=map c [0..<n]" and a20:"ant=andLists (antOfAs@antOfBs)" and a21:"Gp0" shows "cktSat nl (ant \<leadsto> (Is0 out))" proof - have "cktSat nl (ant\<leadsto>cons)" proof (rule caseComparator )qed then have "cktSat nl (ant \<leadsto>cons0)" proof(rule_tac ?B="cons" in steTrans) show "cktSat nl (cons \<leadsto> cons0)" proof(simp only:a16,rule andMono)qed qed with a21 a14 have "Gp0-->cktSat nl (ant \<leadsto> (Is0 out))" proof(rule_tac steGuardI)qed(auto) with a21 show ?thesis by simp qed lemma caseParaComparatorEq: assumes a1:"xorTab=[[ONE,ONE],[ZERO,ZERO]]" and a2:"andTab=[andLine]" and (* a, b, c: three functions map int to nodes a3:"InputOfAs= {l. EX i<n. l=Input (a i)}" and a4:"InputOfBs= {l. EX i<n. l=Input (b i)}" and*) a5:"xorGLs={ l. EX i<n. l= Gate (c i) [(a i),(b i)] xorTab }" and a6:"andLine=map (λ n. ONE) [0..<n]" and a7:"andG=Gate out cs andTab" and a8:" (xorGLs Un {andG} )⊆ nl " and a9:"nl:netlists" and (*bvOfAs, bvOfBs *) a10:"antOfAs=map (λ i. ( Isb (a i) (bvOfAs i))) [0..<n] " and a11:"antOfBs=map (λ i. ( Isb (b i) (bvOfBs i))) [0..<n]" and a12:"Gp0=(EX i. i< n & ( bvOfAs i) ~=( bvOfBs i))" and a13:"Gp1=(ALL i. i< n --> ( bvOfAs i) =( bvOfBs i))" and a14:"cons0=Gp0->T Is0 out " and a15:"cons1=Gp1->T Is1 out " and a16:"cons=cons0 andT cons1" and a17:"isClosed nl" and a18:"0<n" and a19:"cs=map c [0..<n]" and a20:"ant=andLists (antOfAs@antOfBs)" and a21:"Gp1" shows "cktSat nl (ant \<leadsto>(Is1 out))" proof - have "cktSat nl (ant\<leadsto>cons)" proof (rule caseComparator )qed then have "cktSat nl (ant \<leadsto>cons1)" proof(rule_tac ?B="cons" in steTrans) show "cktSat nl (cons \<leadsto> cons1)" proof(simp only:a16) have b1:"cktSat nl ((cons0 andT cons1) \<leadsto> (cons1 andT cons0))" proof(rule steCongCons) show "defSqOfTrForm (cons1 andT cons0) =defSqOfTrForm (cons0 andT cons1) " proof(rule steAndComm)qed show " cktSat nl (cons0 andT cons1 \<leadsto> cons0 andT cons1)" proof(rule steRef) qed qed have "cktSat nl ((cons1 andT cons0) \<leadsto>cons1)" proof(rule andMono)qed with b1 show "cktSat nl ((cons0 andT cons1) \<leadsto>cons1)" proof(rule_tac steTrans)qed qed qed with a15 have "Gp1-->cktSat nl (ant \<leadsto> (Is1 out))" proof(rule_tac steGuardI)qed(auto) with a21 show ?thesis by simp qed end
lemma andMono:
[|nl ∈ netlists; isClosed nl|] ==> cktSat nl (A andT B \<leadsto> A)
lemma andListsAllI1:
[|nl ∈ netlists; isClosed nl|] ==> (∀a. a ∈ set A --> cktSat nl (ant \<leadsto> a)) --> cktSat nl (ant \<leadsto> andLists A)
lemma andListsAllI:
[|nl ∈ netlists; isClosed nl; ∀a. a ∈ set A --> cktSat nl (ant \<leadsto> a)|] ==> cktSat nl (ant \<leadsto> andLists A)
lemma andListMem:
[|nl ∈ netlists; isClosed nl|] ==> a mem A --> cktSat nl (andLists A \<leadsto> andLists [a])
lemma andListMem':
[|nl ∈ netlists; isClosed nl; a mem A|] ==> cktSat nl (andLists A \<leadsto> andLists [a])
lemma andListMem1:
[|nl ∈ netlists; isClosed nl; a mem A|] ==> cktSat nl (andLists A \<leadsto> a)
lemma andListSubsetEq:
[|nl ∈ netlists; isClosed nl; set B ∩ - {chaos} ⊆ set A|] ==> cktSat nl (andLists A \<leadsto> andLists B)
lemma andListCat:
[|nl ∈ netlists; isClosed nl; cktSat nl (andLists A1.0 \<leadsto> andLists B1.0); cktSat nl (andLists A2.0 \<leadsto> andLists B2.0)|] ==> cktSat nl (andLists (A1.0 @ A2.0) \<leadsto> andLists (B1.0 @ B2.0))
lemma caseAux1:
[|xorTab = [[ONE, ONE], [ZERO, ZERO]]; andTab = [andLine]; xorGLs = {l. ∃i<n. l = Gate (c i) [a i, b i] xorTab}; andLine = map (λn. ONE) [0..<n]; andG = Gate out cs andTab; xorGLs ∪ {andG} ⊆ nl; nl ∈ netlists; antOfAs = map (λi. Isb (a i) (bvOfAs i)) [0..<n]; antOfBs = map (λi. Isb (b i) (bvOfBs i)) [0..<n]; i < n ∧ bvOfAs i = bvOfBs i; cons0.0 = Gp0.0 ->T Is0 out; cons1.0 = Gp1.0 ->T Is1 out; cons = cons0.0 andT cons1.0; isClosed nl; 0 < n; cs = map c [0..<n]; ant = andLists (antOfAs @ antOfBs)|] ==> cktSat nl (ant \<leadsto> Is1 (c i))
lemma caseAux1:
[|xorTab = [[ONE, ONE], [ZERO, ZERO]]; andTab = [andLine]; xorGLs = {l. ∃i<n. l = Gate (c i) [a i, b i] xorTab}; andLine = map (λn. ONE) [0..<n]; andG = Gate out cs andTab; xorGLs ∪ {andG} ⊆ nl; nl ∈ netlists; antOfAs = map (λi. Isb (a i) (bvOfAs i)) [0..<n]; antOfBs = map (λi. Isb (b i) (bvOfBs i)) [0..<n]; cons0.0 = Gp0.0 ->T Is0 out; cons1.0 = Gp1.0 ->T Is1 out; cons = cons0.0 andT cons1.0; isClosed nl; 0 < n; cs = map c [0..<n]; ant = andLists (antOfAs @ antOfBs); ∀i<n. bvOfAs i = bvOfBs i|] ==> cktSat nl (ant \<leadsto> andLists (map Is1 cs))
lemma caseComparator:
[|xorTab = [[ONE, ONE], [ZERO, ZERO]]; andTab = [andLine]; xorGLs = {l. ∃i<n. l = Gate (c i) [a i, b i] xorTab}; andLine = map (λn. ONE) [0..<n]; andG = Gate out cs andTab; xorGLs ∪ {andG} ⊆ nl; nl ∈ netlists; antOfAs = map (λi. Isb (a i) (bvOfAs i)) [0..<n]; antOfBs = map (λi. Isb (b i) (bvOfBs i)) [0..<n]; Gp0.0 = (∃i<n. bvOfAs i ≠ bvOfBs i); Gp1.0 = (∀i<n. bvOfAs i = bvOfBs i); cons0.0 = Gp0.0 ->T Is0 out; cons1.0 = Gp1.0 ->T Is1 out; cons = cons0.0 andT cons1.0; isClosed nl; 0 < n; cs = map c [0..<n]; ant = andLists (antOfAs @ antOfBs)|] ==> cktSat nl (ant \<leadsto> cons)
lemma caseParaComparatorNeq:
[|xorTab = [[ONE, ONE], [ZERO, ZERO]]; andTab = [andLine]; xorGLs = {l. ∃i<n. l = Gate (c i) [a i, b i] xorTab}; andLine = map (λn. ONE) [0..<n]; andG = Gate out cs andTab; xorGLs ∪ {andG} ⊆ nl; nl ∈ netlists; antOfAs = map (λi. Isb (a i) (bvOfAs i)) [0..<n]; antOfBs = map (λi. Isb (b i) (bvOfBs i)) [0..<n]; Gp0.0 = (∃i<n. bvOfAs i ≠ bvOfBs i); Gp1.0 = (∀i<n. bvOfAs i = bvOfBs i); cons0.0 = Gp0.0 ->T Is0 out; cons1.0 = Gp1.0 ->T Is1 out; cons = cons0.0 andT cons1.0; isClosed nl; 0 < n; cs = map c [0..<n]; ant = andLists (antOfAs @ antOfBs); Gp0.0|] ==> cktSat nl (ant \<leadsto> Is0 out)
lemma caseParaComparatorEq:
[|xorTab = [[ONE, ONE], [ZERO, ZERO]]; andTab = [andLine]; xorGLs = {l. ∃i<n. l = Gate (c i) [a i, b i] xorTab}; andLine = map (λn. ONE) [0..<n]; andG = Gate out cs andTab; xorGLs ∪ {andG} ⊆ nl; nl ∈ netlists; antOfAs = map (λi. Isb (a i) (bvOfAs i)) [0..<n]; antOfBs = map (λi. Isb (b i) (bvOfBs i)) [0..<n]; Gp0.0 = (∃i<n. bvOfAs i ≠ bvOfBs i); Gp1.0 = (∀i<n. bvOfAs i = bvOfBs i); cons0.0 = Gp0.0 ->T Is0 out; cons1.0 = Gp1.0 ->T Is1 out; cons = cons0.0 andT cons1.0; isClosed nl; 0 < n; cs = map c [0..<n]; ant = andLists (antOfAs @ antOfBs); Gp1.0|] ==> cktSat nl (ant \<leadsto> Is1 out)