theory trajForm=net: datatype trajForm= Is1 node |Is0 node |TAND trajForm trajForm (infixl "andT" 65) |When bool trajForm (infixr "->T" 65) |Next trajForm |chaos consts valid ::"( stateSeq)=>trajForm=>bool" ( "(_ \<Turnstile> _)"[80,80] 80) primrec "asq \<Turnstile> (Is1 n) = T \<sqsubseteq> ( asq 0 n)" "asq \<Turnstile> (Is0 n) = F \<sqsubseteq> ( asq 0 n)" "asq \<Turnstile> (f1 andT f2)= (asq \<Turnstile> f1 & asq \<Turnstile> f2)" "asq \<Turnstile>(P ->T tf) =( P --> asq \<Turnstile> tf)" "asq \<Turnstile>( Next tf) =((suffix 1 asq)\<Turnstile> tf)" "asq \<Turnstile>(chaos) =True" constdefs trajOfCirc::"entity set => stateSeq set" "trajOfCirc enttrs == {seq. fSeq enttrs seq = seq}" datatype assertion=Leadsto trajForm trajForm (infixr "\<leadsto>" 50) consts cktSat ::"entity set => assertion => bool" primrec cktSat_def: "cktSat enttrs (A\<leadsto> C)= ( ∀ tr. tr :(trajOfCirc enttrs ) --> ( tr\<Turnstile> A --> tr\<Turnstile> C ) )" consts defSqOfTrForm ::"trajForm => stateSeq" primrec defSqIs1_def: " defSqOfTrForm (Is1 n)=(λ t m. (if (t=0 & m=n) then T else X))" defSqIs0_def: " defSqOfTrForm (Is0 n)=(λ t m. if (t=0 & m=n) then F else X)" defSqAnd_def: " defSqOfTrForm (A andT B)=(λ t m. lub (defSqOfTrForm A t m) (defSqOfTrForm B t m)) " defSqGuard_def: " defSqOfTrForm (P ->T F0)=(λ t m. let v=(defSqOfTrForm F0 t m) in (P-->(fst v), P--> (snd v))) " defSqNext_def: " defSqOfTrForm (Next F0)=(λ t m. let v=(defSqOfTrForm F0 (t - 1) m) in if (t~=0) then v else X)" defSqChaos_def:" defSqOfTrForm chaos =(λ t m. X)" consts onGuards::" trajForm => bool set" primrec "onGuards (Is1 n) = {}" "onGuards (Is0 n) = {}" "onGuards (f1 andT f2)= (onGuards f1) Un (onGuards f2)" "onGuards (P ->T tf) =(onGuards tf) Un {P}" "onGuards ( Next tf) =(onGuards tf)" "onGuards chaos={}" constdefs defTrajOfCirc::"trajForm => entity set => stateSeq" "defTrajOfCirc trajF enttrs == fSeq enttrs (defSqOfTrForm trajF)" lemma XIsLeastValue: " X \<sqsubseteq> a" apply(unfold leq_def) apply(unfold lub_def) apply(unfold X_def,simp) done lemma lubMeans: "[| x \<sqsubseteq> a; y\<sqsubseteq> a|] ==> lub x y \<sqsubseteq> a" apply(unfold leq_def lub_def) apply auto apply(simp add:Pair_fst_snd_eq) done lemma defSeqIsLeast: shows "ALL seq. seq \<Turnstile> f --> (defSqOfTrForm f)\<sqsubseteq>sq seq " (is "ALL seq. ?P1 seq f --> ?P2 seq f") proof(induct_tac f) fix m let ?f="(Is1 m)" show "ALL seq. ?P1 seq ?f --> ?P2 seq ?f" proof(rule allI,rule impI) fix seq assume a1:"?P1 seq ?f" show "?P2 seq ?f" apply(simp,unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(cut_tac a1) apply(simp add:XIsLeastValue) done qed next fix m let ?f="(Is0 m)" show "ALL seq. ?P1 seq ?f --> ?P2 seq ?f" proof(rule allI,rule impI) fix seq assume a1:"?P1 seq ?f" show "?P2 seq ?f" apply(simp,unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(cut_tac a1) apply(simp add:XIsLeastValue) done qed next fix trajForm1 trajForm2 assume a1:"ALL seq. ?P1 seq trajForm1--> ?P2 seq trajForm1 " and a2:"ALL seq. ?P1 seq trajForm2--> ?P2 seq trajForm2" let ?f="(trajForm1 andT trajForm2)" show "ALL seq. ?P1 seq ?f --> ?P2 seq ?f" proof(rule allI,rule impI) fix seq assume b1:"?P1 seq ?f" show "?P2 seq ?f" proof - have c1:" defSqOfTrForm trajForm1 \<sqsubseteq>sq seq" apply(cut_tac a1 b1) by auto have c2:" defSqOfTrForm trajForm2 \<sqsubseteq>sq seq" apply(cut_tac a2 b1) by auto show ?thesis apply(simp,unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(cut_tac c1 c2) apply(unfold stateSqLeq_def stateLeq_def) apply(rule lubMeans) apply auto done qed qed next fix P trajF assume a1:"ALL seq. ?P1 seq trajF --> ?P2 seq trajF" let ?f="P ->T trajF" show "ALL seq. ?P1 seq ?f --> ?P2 seq ?f" apply(rule allI) apply(simp,unfold stateSqLeq_def stateLeq_def) apply(rule impI) apply(rule allI)+ apply(cut_tac a1 ) apply(unfold stateSqLeq_def stateLeq_def) apply(case_tac "P=True") apply simp apply(drule_tac x="seq" in spec) apply simp apply(drule_tac x="t" in spec) apply(drule_tac x="n" in spec) apply( simp add:Let_def) apply simp apply(fold X_def) apply(simp add:XIsLeastValue) done next fix trajForm assume a1:"ALL seq. ?P1 seq trajForm -->?P2 seq trajForm" let ?f="(Next trajForm)" show "ALL seq. ?P1 seq ?f --> ?P2 seq ?f" apply(rule allI) apply(simp,unfold stateSqLeq_def stateLeq_def) apply(rule impI) apply(rule allI)+ apply(cut_tac a1 ) apply(drule_tac x="suffix (Suc 0) seq" in spec) apply(unfold stateSqLeq_def stateLeq_def) apply(case_tac "0 < t",simp) apply(simp add:Let_def) apply(drule_tac x="(t - Suc 0)" in spec) apply(unfold suffix_def) apply auto apply(rule XIsLeastValue) done next let ?f=" chaos" show "ALL seq. ?P1 seq ?f --> ?P2 seq ?f" apply(rule allI) apply(simp,unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(rule XIsLeastValue) done qed lemma defSeqIsLeast: shows "[| seq \<Turnstile> f |]==> (defSqOfTrForm f)\<sqsubseteq>sq seq " apply(cut_tac defSeqIsLeast[where f="f"]) by blast lemma satIsMono: "ALL sq sq1. sq \<Turnstile> f --> sq \<sqsubseteq>sq sq1 --> sq1 \<Turnstile> f" apply(induct_tac f) apply(unfold stateSqLeq_def stateLeq_def lub_def) apply(auto dest:leqIsTrans) apply(drule_tac x="suffix (Suc 0) sq" in spec) apply simp apply(drule_tac x="suffix (Suc 0) sq1" in spec) apply(unfold suffix_def,auto dest:leqIsTrans) done lemma satIsMono: "[|sq \<Turnstile> f ; sq \<sqsubseteq>sq sq1 |] ==> sq1 \<Turnstile> f" apply(cut_tac satIsMono[where f="f"]) apply blast done lemma lubSelf: " a \<sqsubseteq> (lub a b )" apply(unfold leq_def lub_def) apply auto done lemma lubId:" lub a a=a" by(unfold lub_def, auto) lemma lubComm:" lub a b= lub b a" apply(unfold leq_def lub_def) apply auto done lemma fclosureInput: assumes a1:"nl:netlists" and a2:"Input x ∈ nl " shows "fclosure nl s x= s x" proof - let ?pair="(x, s x)" have b1:"?pair :rclosure nl s " apply(rule stAddInput) by assumption have b2:"isDefinedIn x nl" apply(unfold isDefinedIn_def) apply(cut_tac a1) apply(rule_tac x="Input x" in bexI) apply simp by assumption then show ?thesis proof(unfold fclosure_def,simp add:Let_def) have "(EX! pair. pair:rclosure nl s & (fst pair)=x) " (is "EX! pair. ?P2 pair") apply(rule nodeNameisUniInrclosure) apply assumption+ apply(cut_tac b2) by assumption then obtain enttr2 where c1: "?P2 enttr2 &(∀ enttr'. ?P2 enttr' --> enttr' =enttr2)" apply - apply(erule_tac P="?P2" in ex1E) apply blast done from this and b1 have " ?pair=enttr2" apply - apply(erule conjE)+ apply(drule_tac x="(x, s x)" in spec) apply auto done then show "snd (THE pair0. pair0 ∈ rclosure nl s ∧ fst pair0 = x) = s x" apply(rule_tac a="enttr2" in theI2) apply(cut_tac c1,simp) apply(cut_tac c1,blast) apply(cut_tac c1) by auto qed qed lemma fclosureDelay: assumes a1:"nl:netlists" and a2:"Delay x data ∈ nl " shows "fclosure nl s x= s x" proof - let ?pair="(x, s x)" have b1:"?pair :rclosure nl s " apply(rule stAddDelay) by assumption have b2:"isDefinedIn x nl" apply(unfold isDefinedIn_def) apply(cut_tac a1) apply(rule_tac x="Delay x data" in bexI) apply simp by assumption then show ?thesis proof(unfold fclosure_def,simp add:Let_def) have "(EX! pair. pair:rclosure nl s & (fst pair)=x) " (is "EX! pair. ?P2 pair") apply(rule nodeNameisUniInrclosure) apply assumption+ apply(cut_tac b2) by assumption then obtain enttr2 where c1: "?P2 enttr2 &(∀ enttr'. ?P2 enttr' --> enttr' =enttr2)" apply - apply(erule_tac P="?P2" in ex1E) apply blast done from this and b1 have " ?pair=enttr2" apply - apply(erule conjE)+ apply(drule_tac x="(x, s x)" in spec) apply auto done then show "snd (THE pair0. pair0 ∈ rclosure nl s ∧ fst pair0 = x) = s x" apply(rule_tac a="enttr2" in theI2) apply(cut_tac c1,simp) apply(cut_tac c1,blast) apply(cut_tac c1) by auto qed qed lemma fclosureGate: assumes a1:"nl:netlists" and a2:"Gate n inps tab ∈ nl " and a3:"length stateLs = length inps" and a4:"∀l. l ∈ set tab --> length l = length inps" and a5:"∀pair0. pair0 mem zip inps stateLs --> pair0 ∈ rclosure nl s" shows "fclosure nl s n= lub (funOfTab tab stateLs) ( s n)" proof - let ?pair="(n, lub (funOfTab tab stateLs) (s n))" have b1:"?pair :rclosure nl s " apply(rule stAddGate) by assumption+ have b2:"isDefinedIn n nl" apply(unfold isDefinedIn_def) apply(cut_tac a1) apply(rule_tac x="Gate n inps tab" in bexI) apply simp by assumption then show ?thesis proof(unfold fclosure_def,simp add:Let_def) have "(EX! pair. pair:rclosure nl s & (fst pair)=n) " (is "EX! pair. ?P2 pair") apply(rule nodeNameisUniInrclosure) apply assumption+ apply(cut_tac b2) by assumption then obtain enttr2 where c1: "?P2 enttr2 &(∀ enttr'. ?P2 enttr' --> enttr' =enttr2)" apply - apply(erule_tac P="?P2" in ex1E) apply blast done from this and b1 have " ?pair=enttr2" apply - apply(erule conjE)+ apply(drule_tac x="?pair" in spec) apply auto done then show "snd (THE pair0. pair0 ∈ rclosure nl s ∧ fst pair0 = n) = lub (funOfTab tab stateLs) ( s n) " apply(rule_tac a="enttr2" in theI2) apply(cut_tac c1,simp) apply(cut_tac c1,blast) apply(cut_tac c1) apply(erule conjE)+ apply(drule_tac x="x" in spec) by auto qed qed lemma lubJieHe: "lub a (lub b c)=lub (lub a b) c" apply(unfold lub_def,auto) done lemma fSeqIsIdAux1: assumes a1:"nl:netlists" and a2:" isDefinedIn n nl" and a3:"pair ∈ rclosure nl (fclosure nl s)" and a4:"( fst pair) = n" shows "pair ∈ (rclosure nl s) " (is "?P pair") using a3 proof(induct) fix x assume b1:"Input x ∈ nl" show "?P (x, fclosure nl s x)" proof - have "fclosure nl s x = s x" by(rule fclosureInput,assumption+) then show ?thesis proof(simp) show "(x, s x) ∈ rclosure nl s" by(rule stAddInput) qed qed next fix data n assume b1:" Delay n data ∈ nl" show "?P (n, fclosure nl s n)" proof - have "fclosure nl s n = s n" by(rule fclosureDelay,assumption+) then show ?thesis proof(simp) show "(n, s n) ∈ rclosure nl s" by(rule stAddDelay) qed qed next fix inps n stateLs tab assume b1:"Gate n inps tab ∈ nl " and b2:"length stateLs = length inps" and b3:"∀l. l ∈ set tab --> length l = length inps" and b4:"∀pair0. pair0 mem zip inps stateLs --> pair0 ∈ rclosure nl (fclosure nl s) ∧ pair0 ∈ rclosure nl s" let ?pair="(n, lub (funOfTab tab stateLs) (fclosure nl s n))" show "?P ?pair" proof - have "fclosure nl s n =lub (funOfTab tab stateLs) ( s n) " apply(rule fclosureGate,assumption+) apply(cut_tac b4,blast) done then show ?thesis proof(simp) show "(n, lub (funOfTab tab stateLs) (lub (funOfTab tab stateLs) (s n)) )∈ rclosure nl s" proof - have "lub (funOfTab tab stateLs) (lub (funOfTab tab stateLs) (s n)) = lub (funOfTab tab stateLs) (s n)" by (simp add:lubJieHe lubId) with this show ?thesis apply - apply( simp,rule stAddGate) apply(assumption+) apply(cut_tac b4,blast) done qed qed qed qed lemma fclosureIsrclosure: "[|nl:netlists; isDefinedIn n nl |] ==> (n,fclosure nl s n): rclosure nl s" apply(unfold fclosure_def) apply(simp add:Let_def) apply(subgoal_tac "(EX! pair. pair:rclosure nl s & (fst pair)=n) " ) apply(erule ex1E) apply(rule_tac a="pair" in theI2) apply blast apply blast apply(subgoal_tac "(n, snd x)=x", simp) apply(simp add:Pair_fst_snd_eq) apply(rule nodeNameisUniInrclosure) by assumption+ lemma fclosureVal: "[|nl:netlists; isDefinedIn n nl;pair:rclosure nl s;fst pair =n |] ==> (fclosure nl s n)=snd pair" apply(unfold fclosure_def) apply(simp add:Let_def) apply(subgoal_tac "(EX! pair. pair:rclosure nl s & (fst pair)=n) " ) apply(erule ex1E) apply(subgoal_tac "paira=pair") apply(rule_tac a="pair" in theI2) apply simp apply blast apply blast apply (drule_tac x="pair" in spec) apply simp apply(rule nodeNameisUniInrclosure) by assumption+ lemma fclosureIsId: assumes a1:"nl:netlists" and a2:"isClosed nl" shows " fclosure nl (fclosure nl s) n = (fclosure nl s) n " (is " ?LHS s n= ?RHS s n") proof(case_tac "isDefinedIn n nl") let ?s="(fclosure nl s)" assume b1:"isDefinedIn n nl" have "(n,fclosure nl ?s n) :rclosure nl ?s" apply(rule fclosureIsrclosure) by assumption+ then have "(n,fclosure nl ?s n): rclosure nl s" thm fSeqIsIdAux1 apply(rule_tac pair="(n,fclosure nl ?s n)" in fSeqIsIdAux1) apply assumption+ by simp then have "fclosure nl s n=fclosure nl ?s n" apply(subgoal_tac "fst (n,fclosure nl ?s n)= n") apply(cut_tac a1 b1 ) apply(drule fclosureVal) apply simp+ done then show ?thesis by simp next assume b1:"¬ isDefinedIn n nl" show "?LHS s n= ?RHS s n" apply(cut_tac b1,unfold fclosure_def) apply auto done qed lemma fSeqNext1: assumes a1:"nl:netlists" and a2:"isClosed nl" and a4:"isDelayNames nl n" shows "(fSeq nl sq (Suc t) n)= (let l= (lookUp nl n) in let inps=fanins l in lub (sq (Suc t) n) ((fSeq nl sq t) (hd inps)))" proof - 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)" have "(fSeq nl sq (Suc t) n)=fclosure nl ?s n" apply(simp add:Let_def) done also have "…=?s n" apply(cut_tac a4) apply(unfold isDelayNames_def) apply(erule exE)+ apply(rule_tac data="data" in fclosureDelay) apply assumption apply auto done also have "…= (let l= (lookUp nl n) in let inps=fanins l in lub (sq (Suc t) n) ((fSeq nl sq t) (hd inps)))" apply(cut_tac a4,simp add:Let_def) done ultimately show ?thesis by auto qed lemma fSeqNext2: assumes a1:"nl:netlists" and a2:"isClosed nl" and a4:"~isDelayNames nl n" and a5:"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)" shows "(fSeq nl sq (Suc t) n)= fclosure nl s n " proof - 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)" have "(fSeq nl sq (Suc t) n)=fclosure nl ?s n" apply(simp add:Let_def) done then show ?thesis apply(cut_tac a5,simp) done qed lemma fSeqIsId: assumes a1:"nl:netlists" and a2:"isClosed nl" shows "ALL n. fSeq nl (fSeq nl sq) t n = fSeq nl sq t n" (is "ALL n. ?P t n") proof(induct_tac t) show "ALL n. ?P 0 n" proof fix n show "fSeq nl (fSeq nl sq) 0 n = fSeq nl sq 0 n" by(simp,rule fclosureIsId) qed next fix t assume b1:"ALL n. ?P t n" show "ALL n. ?P (Suc t) n" proof(rule allI) fix n let ?sq=" (fSeq nl sq)" let ?Ps1=" λ sq x . if ( isDelayNames nl x) then ( let l= (lookUp nl x) in let inps=fanins l in let v1=lub (sq (Suc t) x) ((fSeq nl sq t) (hd inps)) in v1 ) else (sq (Suc t)) x " let ?s1=" ?Ps1 ?sq" let ?inp="(hd (fanins ( (lookUp nl n))))" let ?s0=" ?Ps1 sq" have b2:"?s1 = fclosure nl ?s0 " proof(rule ext) fix na let ?inpa="(hd (fanins ( (lookUp nl na))))" from b1 have d2:"fSeq nl ?sq t ?inpa = fSeq nl sq t ?inpa" apply - by blast have "isDelayNames nl na | ~isDelayNames nl na " by simp moreover {assume c1:"isDelayNames nl na" have d5:"isDefinedIn na nl" apply(cut_tac c1 a1) apply(rule latchNameInNet) by assumption+ have c3:"?sq (Suc t) na= (let l= (lookUp nl na) in let inps=fanins l in lub (sq (Suc t) na) ((fSeq nl sq t) (hd inps)))" apply(rule fSeqNext1) apply assumption+ done have "?s1 na= lub (fSeq nl sq (Suc t) na) (fSeq nl (fSeq nl sq) t ?inpa)" apply(cut_tac c1, simp add:Let_def) done also have "…=lub (fSeq nl sq (Suc t) na) (fSeq nl sq t ?inpa)" by(cut_tac d2,simp) also have "…=lub (lub (sq (Suc t) na) ((fSeq nl sq t) ?inpa )) ((fSeq nl sq t) ?inpa)" apply(cut_tac c3,simp add:Let_def) done also have "…=lub (sq (Suc t) na) (lub (fSeq nl sq t ?inpa) ((fSeq nl sq t) ?inpa )) " apply(simp add:lubJieHe[THEN sym]) done also have "… = lub (sq (Suc t) na) ((fSeq nl sq t) ?inpa )" thm lubSelf apply (simp add:lubId) done also have "…=fclosure nl ?s0 na" apply(rule sym) apply(subgoal_tac "fclosure nl ?s0 na = ?s0 na") apply(cut_tac c1, simp add:Let_def) apply(cut_tac c1) apply(unfold isDelayNames_def) apply(erule exE)+ apply(rule_tac data="data" in fclosureDelay) apply assumption apply auto done finally have "?s1 na==fclosure nl ?s0 na" by auto } moreover {assume c1:"~isDelayNames nl na" have "?s1 na= (fSeq nl sq) (Suc t) na" by(cut_tac c1,simp) also have "…=fclosure nl ?s0 na" apply( simp add:Let_def) done finally have "?s1 na=fclosure nl ?s0 na" by auto } ultimately show "?s1 na = fclosure nl ?s0 na " by blast qed then show " ?P (Suc t) n" proof - have "fSeq nl (fSeq nl sq) (Suc t) n = fclosure nl ?s1 n" apply (simp add:Let_def) done also have "…=fclosure nl (fclosure nl ?s0) n" apply(cut_tac b2,simp) done also have "…=(fclosure nl ?s0) n" apply(rule fclosureIsId) apply assumption+ done also have "…=fSeq nl sq (Suc t) n" apply(rule sym) apply(simp add:Let_def) done finally show ?thesis by auto qed qed qed lemma fSeqIsId: assumes a1:"nl:netlists" and a2:"isClosed nl" shows " fSeq nl (fSeq nl sq) = fSeq nl sq " apply(rule ext)+ apply(subgoal_tac "ALL n. fSeq nl (fSeq nl sq) x n = fSeq nl sq x n") apply blast apply(rule fSeqIsId) apply assumption+ done lemma defTrajIsTraj: " [|nl:netlists ; isClosed nl|] ==> (defTrajOfCirc f nl):trajOfCirc nl " apply(unfold trajOfCirc_def) apply simp apply(unfold defTrajOfCirc_def) apply(rule fSeqIsId) apply assumption+ done lemma allSeqGtDefSeqSatF:"(defSqOfTrForm f) \<Turnstile> f" (is "?P f") proof(induct_tac f) fix m let ?f="(Is1 m)" show "?P ?f" apply auto by (rule leqReflexive) next fix m let ?f="(Is0 m)" show "?P ?f" apply auto by (rule leqReflexive) next fix trajForm1 trajForm2 assume a1:"?P trajForm1 " and a2:"?P trajForm2" let ?f="(trajForm1 andT trajForm2)" show "?P ?f" apply(cut_tac a1 a2) apply simp apply(rule conjI) apply(subgoal_tac "defSqOfTrForm trajForm1 \<sqsubseteq>sq (λu ua. lub (defSqOfTrForm trajForm1 u ua) (defSqOfTrForm trajForm2 u ua))") apply(rule satIsMono) apply simp+ apply(unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(auto intro:lubSelf) apply(rule satIsMono) apply simp+ apply(unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(subgoal_tac "lub (defSqOfTrForm trajForm1 t n) (defSqOfTrForm trajForm2 t n)= lub (defSqOfTrForm trajForm2 t n) (defSqOfTrForm trajForm1 t n)") apply(auto intro:lubSelf) apply(simp add: lubComm) done next fix P trajF assume a1:"?P trajF " let ?f=" (P ->T trajF)" show "?P ?f" apply(cut_tac a1 ) apply simp apply(case_tac "P=True") apply(subgoal_tac "(λu ua. let v = defSqOfTrForm trajF u ua in v) = defSqOfTrForm trajF") apply simp apply(simp add:Let_def) apply simp done next fix trajF assume a1:"?P trajF " let ?f=" (Next trajF)" show "?P ?f" apply simp apply(unfold suffix_def,simp) apply(subgoal_tac "(λt u. let v = defSqOfTrForm trajF t u in v) = defSqOfTrForm trajF") apply simp apply(cut_tac a1) apply(simp add:Let_def)+ done next let ?f=" chaos" show "?P ?f" apply simp done qed lemma allSeqGtDefSeqSatF: " defSqOfTrForm f \<sqsubseteq>sq seq ==> seq \<Turnstile> f" apply(cut_tac allSeqGtDefSeqSatF[where f="f"]) apply(blast dest:satIsMono) done lemma defTrajIsLeast: "[|nl:netlists; isClosed nl; τ : trajOfCirc nl; τ \<Turnstile> f; isDefinedIn n nl|]==> (defTrajOfCirc f nl ) t n \<sqsubseteq> τ t n" apply(subgoal_tac "(defSqOfTrForm f)\<sqsubseteq>sq τ ") thm defSeqIsLeast prefer 2 apply(blast dest:defSeqIsLeast) apply(unfold defTrajOfCirc_def) apply(subgoal_tac "fSeq nl (defSqOfTrForm f) t n \<sqsubseteq> fSeq nl τ t n") thm fSeqIsMono prefer 2 apply(drule_tac ?sq1.0=" defSqOfTrForm f" and ?sq2.0=" τ" in fSeqIsMono) apply assumption+ apply(unfold trajOfCirc_def) apply simp done lemma defTrajIsLeast: "[|nl:netlists; isClosed nl; τ : trajOfCirc nl; τ \<Turnstile> f|]==> (defTrajOfCirc f nl )\<sqsubseteq>sq τ " apply(unfold stateSqLeq_def stateLeq_def ) apply(rule allI)+ apply(case_tac "isDefinedIn n nl") apply(drule_tac t="t" and n="n" in defTrajIsLeast) apply assumption+ apply(unfold defTrajOfCirc_def ) apply(induct_tac t) apply simp apply(subgoal_tac "(fclosure nl (defSqOfTrForm f 0) n) =(defSqOfTrForm f 0) n") apply simp apply(cut_tac defSeqIsLeast[where seq=" τ" and f="f" ]) apply(unfold stateSqLeq_def stateLeq_def,simp+) apply(unfold fclosure_def) apply simp apply(subgoal_tac "fSeq nl (defSqOfTrForm f) (Suc na) n =(defSqOfTrForm f (Suc na)) n") apply( simp del:fSeqn) apply(cut_tac defSeqIsLeast[where seq=" τ" and f="f" ]) apply(unfold stateSqLeq_def stateLeq_def) apply(drule_tac x="Suc na" in spec) apply(drule_tac x="n" in spec) apply blast+ apply(simp add:Let_def) apply(subgoal_tac "~isDelayNames nl n") apply(unfold fclosure_def, auto) apply(blast dest:latchNameInNet) done lemma sqLeqIsTrans: "[|sq0 \<sqsubseteq>sq sq1; sq1 \<sqsubseteq>sq sq2 |]==> sq0 \<sqsubseteq>sq sq2" apply(unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(blast intro:leqIsTrans) done lemma allTrajGtDefTrajSatF: "[|nl:netlists; isClosed nl; τ : trajOfCirc nl; (defTrajOfCirc f nl ) \<sqsubseteq>sq τ |]==> τ \<Turnstile> f" apply(subgoal_tac "defSqOfTrForm f \<sqsubseteq>sq defTrajOfCirc f nl" ) prefer 2 apply(unfold defTrajOfCirc_def) thm fSeqIsMono apply(rule fSeqIsExtensive) apply assumption+ apply(rule allSeqGtDefSeqSatF) apply(blast dest:sqLeqIsTrans) done lemma fundmental1: assumes a1:"nl:netlists" and a2:"(defSqOfTrForm C ) \<sqsubseteq>sq(defTrajOfCirc A nl ) " and a3:"isClosed nl" shows "cktSat nl ( A \<leadsto> C)" proof(simp, rule allI,rule impI) fix tr assume a1:"tr ∈ trajOfCirc nl" show " tr \<Turnstile> A --> tr \<Turnstile> C" proof assume b1:"tr \<Turnstile> A" show "tr \<Turnstile> C" proof(rule allSeqGtDefSeqSatF) show "defSqOfTrForm C \<sqsubseteq>sq tr" proof - from a1 b1 have "defTrajOfCirc A nl \<sqsubseteq>sq tr" apply - thm allTrajGtDefTrajSatF apply(rule defTrajIsLeast) apply assumption+ done with a2 show ?thesis by (blast dest: sqLeqIsTrans) qed qed qed qed lemma sqLeqIsReflexive:" sq \<sqsubseteq>sq sq" apply(unfold stateSqLeq_def stateLeq_def) apply(rule allI)+ apply(rule leqReflexive) done lemma fundmental2: assumes a1:"nl:netlists" and a2:"cktSat nl ( A \<leadsto> C) " and a3:"isClosed nl" shows "(defSqOfTrForm C ) \<sqsubseteq>sq(defTrajOfCirc A nl )" proof - have b1:"(defTrajOfCirc A nl ):(trajOfCirc nl )" apply(rule defTrajIsTraj) by assumption+ have "(defTrajOfCirc A nl ) \<Turnstile> A" proof(rule allTrajGtDefTrajSatF, (assumption)+, rule defTrajIsTraj, (assumption )+, rule sqLeqIsReflexive) qed with a2 b1 have "(defTrajOfCirc A nl ) \<Turnstile> C" apply - apply(simp) done then show ?thesis by(rule defSeqIsLeast) qed end
lemma XIsLeastValue:
X \<sqsubseteq> a
lemma lubMeans:
[|x \<sqsubseteq> a; y \<sqsubseteq> a|] ==> lub x y \<sqsubseteq> a
lemma defSeqIsLeast:
∀seq. seq \<Turnstile> f --> defSqOfTrForm f \<sqsubseteq>sq seq
lemma defSeqIsLeast:
seq \<Turnstile> f ==> defSqOfTrForm f \<sqsubseteq>sq seq
lemma satIsMono:
∀sq sq1. sq \<Turnstile> f --> sq \<sqsubseteq>sq sq1 --> sq1 \<Turnstile> f
lemma satIsMono:
[|sq \<Turnstile> f; sq \<sqsubseteq>sq sq1.0|] ==> sq1.0 \<Turnstile> f
lemma lubSelf:
a \<sqsubseteq> lub a b
lemma lubId:
lub a a = a
lemma lubComm:
lub a b = lub b a
lemma fclosureInput:
[|nl ∈ netlists; Input x ∈ nl|] ==> fclosure nl s x = s x
lemma fclosureDelay:
[|nl ∈ netlists; Delay x data ∈ nl|] ==> fclosure nl s x = s x
lemma fclosureGate:
[|nl ∈ netlists; Gate n inps tab ∈ nl; length stateLs = length inps; ∀l. l ∈ set tab --> length l = length inps; ∀pair0. pair0 mem zip inps stateLs --> pair0 ∈ rclosure nl s|] ==> fclosure nl s n = lub (funOfTab tab stateLs) (s n)
lemma lubJieHe:
lub a (lub b c) = lub (lub a b) c
lemma fSeqIsIdAux1:
[|nl ∈ netlists; isDefinedIn n nl; pair ∈ rclosure nl (fclosure nl s); fst pair = n|] ==> pair ∈ rclosure nl s
lemma fclosureIsrclosure:
[|nl ∈ netlists; isDefinedIn n nl|] ==> (n, fclosure nl s n) ∈ rclosure nl s
lemma fclosureVal:
[|nl ∈ netlists; isDefinedIn n nl; pair ∈ rclosure nl s; fst pair = n|] ==> fclosure nl s n = snd pair
lemma fclosureIsId:
[|nl ∈ netlists; isClosed nl|] ==> fclosure nl (fclosure nl s) n = fclosure nl s n
lemma fSeqNext1:
[|nl ∈ netlists; isClosed nl; isDelayNames nl n|] ==> fSeq nl sq (Suc t) n = (let l = lookUp nl n; inps = fanins l in lub (sq (Suc t) n) (fSeq nl sq t (hd inps)))
lemma fSeqNext2:
[|nl ∈ netlists; isClosed nl; ¬ isDelayNames nl n; 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)|] ==> fSeq nl sq (Suc t) n = fclosure nl s n
lemma fSeqIsId:
[|nl ∈ netlists; isClosed nl|] ==> ∀n. fSeq nl (fSeq nl sq) t n = fSeq nl sq t n
lemma fSeqIsId:
[|nl ∈ netlists; isClosed nl|] ==> fSeq nl (fSeq nl sq) = fSeq nl sq
lemma defTrajIsTraj:
[|nl ∈ netlists; isClosed nl|] ==> defTrajOfCirc f nl ∈ trajOfCirc nl
lemma allSeqGtDefSeqSatF:
defSqOfTrForm f \<Turnstile> f
lemma allSeqGtDefSeqSatF:
defSqOfTrForm f \<sqsubseteq>sq seq ==> seq \<Turnstile> f
lemma defTrajIsLeast:
[|nl ∈ netlists; isClosed nl; τ ∈ trajOfCirc nl; τ \<Turnstile> f; isDefinedIn n nl|] ==> defTrajOfCirc f nl t n \<sqsubseteq> τ t n
lemma defTrajIsLeast:
[|nl ∈ netlists; isClosed nl; τ ∈ trajOfCirc nl; τ \<Turnstile> f|] ==> defTrajOfCirc f nl \<sqsubseteq>sq τ
lemma sqLeqIsTrans:
[|sq0.0 \<sqsubseteq>sq sq1.0; sq1.0 \<sqsubseteq>sq sq2.0|] ==> sq0.0 \<sqsubseteq>sq sq2.0
lemma allTrajGtDefTrajSatF:
[|nl ∈ netlists; isClosed nl; τ ∈ trajOfCirc nl; defTrajOfCirc f nl \<sqsubseteq>sq τ|] ==> τ \<Turnstile> f
lemma fundmental1:
[|nl ∈ netlists; defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc A nl; isClosed nl|] ==> cktSat nl (A \<leadsto> C)
lemma sqLeqIsReflexive:
sq \<sqsubseteq>sq sq
lemma fundmental2:
[|nl ∈ netlists; cktSat nl (A \<leadsto> C); isClosed nl|] ==> defSqOfTrForm C \<sqsubseteq>sq defTrajOfCirc A nl