Theory net

Up to index of Isabelle/HOL/net

theory net
imports valuePart
begin

theory net=valuePart:

datatype entity =
  Input node |
  Gate  node "node list" PLA|
  Delay node node 

consts isAndGate::"entity => bool"
primrec 
  "isAndGate (Input x)=False"
  "isAndGate ( Delay out x)=False"
  "isAndGate (Gate out inps pla)= isAndTab pla"

consts fanins ::"entity=> (node list)"
primrec 
  "fanins (Input x)=[]"
  "fanins (Gate out inps pla)=inps"
  "fanins (Delay out data )=[data]"

consts fanout:: "entity=> (node)"
 primrec 
   fanoutInput:"fanout (Input x)=x"
   fanoutGate:"fanout (Gate out inps pla)=out"
   fanoutDelay: "fanout (Delay out data) =out"

consts names:: "entity=> node set"
 primrec 
  "names (Input x)={x}"
  "names (Gate out inps pla)={out} Un set inps"
  "names (Delay out data  )={out,data}"

consts faninsOfSet::"entity set => node set"   



types trans="state=>boolPairs"
      entityTransPair="entity× trans"

constdefs entityPart::" entityTransPair => entity"
"entityPart x==fst x"

constdefs transPart::" entityTransPair => trans"
"transPart x== (snd x)"

constdefs isDefinedIn ::"node =>entity set =>bool"
 "isDefinedIn n ens ==EX l:ens. fanout ( l)=n"

consts entitiesOf ::" entityTransPair set =>entity set"

constdefs funOf ::" entityTransPair set  =>state => node =>boolPairs"
 "funOf ens s n==
              let et=THE enttr. (enttr):ens 
                             &((fst enttr)= Input n |
                               ( EX data.(fst enttr)=Delay n data)|
                               (EX inps tab.(fst enttr)= Gate n inps tab)) in
                     (transPart et) s"

(*constdefs nodes ::" entity set => node set"
"nodes enttrs=={n. EX l. (l:enttrs & fanout ( l)=n)}"*)

constdefs lookUp::"entity set =>node =>entity"
"lookUp enttrs n==
              let et=THE enttr. (enttr):enttrs 
                             &((enttr)= Input n |
                               ( EX data  .(enttr)=Delay n data)|
                               (EX inps tab.( enttr)= Gate n inps tab)) in
                     ( et)"
 

consts netlists :: "(entity set) set"
inductive "netlists"
intros 
nilNetList: 
 "{}:netlists "
addInput: 
  "[|¬ isDefinedIn n nl;
    nl :netlists |]==> {(Input n)} Un nl:netlists"
addDelay:
  "[|¬ isDefinedIn n nl; 
    nl :netlists|]==>
 {((Delay n data ))} Un nl:netlists"

addGate :
   "[|(¬ isDefinedIn n nl)  ; ∀ l. (l∈ (set tab)) --> length l=length inps;
      ∀n. n mem inps -->( isDefinedIn n nl);
      nl:netlists |]==> {(Gate n inps tab) } Un nl:netlists"


thm netlists.induct
thm netlists.intros
thm netlists.elims

types
 nodeFunType="node=>node"
 nodeSetType="node set"

constdefs namesOfNet::"(entityTransPair set)=> node set"
" namesOfNet nl =={n. EX l. l:nl & (n:names (fst l))}"

constdefs isClosed::"(entity set)=>bool"
"isClosed  nl == ALL m n. isDefinedIn m nl-->n:set (fanins ((lookUp nl m))) --> isDefinedIn n nl"

(*constdefs sym::"nodeFunType => nodeSetType =>nodeSetType =>(entityTransPair set)=>bool"
  "sym f M  N nt== bij f & ((M Un N) ⊆ nodes nt) &
                  (∀ n. n:N-->( EX m. m: M &f m=n))&
                  (∀ m. m:M -->                  
                   (let lx=entityPart (lookUp nt m) in
                   let ly=entityPart (lookUp nt (f m)) in
                  (case ( lx)  of
                   Input x => ly= Input ( f x)|
                   Delay out data en set0  reset lt => 
                        (
                        (ly= Delay (f out) (f data) (f en) (f set0)  (f reset) lt))|
                   Gate out inps tab => 
                        ( ly= Gate (f out) (map f inps) tab))))" *)



(*constdefs nodeSym::"nodeFunType =>node=>node=>(entityTransPair set)=>bool"
  "nodeSym f m n nt==
                   let fanInDfsOfm=fanInDfsPart (lookUp nt m) in 
                   let fanInDfsOfn=fanInDfsPart (lookUp nt n) in
                   sym f  fanInDfsOfm    fanInDfsOfn  nt " *)

consts rclosure::"entity set => state => (node× boolPairs) set"
inductive " rclosure nl s "
intros
  stAddInput: "[| Input x:nl|]==> (x,s x):rclosure nl s"
  stAddDelay:"[|Delay n data :nl|]==>(n, s n):rclosure nl s"
  stAddGate:"[|Gate n inps tab:nl;
             length stateLs =length inps;
              ∀ l. (l∈ (set tab)) --> length l=length inps;
             ALL pair0. pair0 mem (zip inps stateLs) --> pair0:rclosure nl s|]==>
            (n,lub (funOfTab tab stateLs)  (s n)):rclosure nl s"

constdefs
fclosure::"entity set => state =>state"
"fclosure enttrs s n== if (isDefinedIn n enttrs)
                       then (let et=THE pair0. (pair0):  (rclosure enttrs  s)
                             &(fst pair0) =n in
                              snd  ( et))
                       else s(n)"

constdefs 
isDelayNames::"entity set =>node =>bool"
"isDelayNames enttrs n==EX l  data  .
  ((Delay n data )=( l) & l:enttrs)"



consts
fSeq::"entity set => stateSeq =>stateSeq"
primrec
fSeq0:"fSeq  enttrs sq 0 = fclosure enttrs (sq 0)"
fSeqn:"fSeq enttrs sq (Suc t) = (let s1=
                          λ x. if ( isDelayNames enttrs x) 
                               then 
                                  ( let l= (lookUp enttrs x) in
                                    let inps=fanins l in
                                    let v1=lub (sq (Suc t) x) ((fSeq  enttrs sq t) (hd inps)) in
                                      v1 )
                                else (sq (Suc t)) x  in
                          fclosure enttrs  s1)"


lemma nodeNameIsUniqueInNet:
assumes a1:"nl:netlists"
shows "ALL n.(  isDefinedIn n nl  -->
 ( EX! enttr.(enttr):nl 
                             &((enttr)= Input n |
                               ( EX data .(enttr)=Delay n data )|
                               (EX inps tab.( enttr)= Gate n inps tab))))"
  (is "ALL n. ( ?P2 n nl--> ?P4 n nl)") 
  using a1
proof(induct)   
  show "ALL n. ( ?P2 n {}--> ?P4 n {})"
    by(simp add:  isDefinedIn_def )
next  
  fix m nl
  assume b1:"¬ isDefinedIn m nl" and
  b2:" nl ∈ netlists" and
  b3:"ALL m. (?P2 m nl--> ?P4 m nl)"
   let ?nl="{(Input m)} ∪ nl" 
  show "ALL ma.( ?P2 ma ?nl--> ?P4 ma ?nl)"
  proof(rule allI,(rule impI)+)
    fix ma
    assume 
    c2:"isDefinedIn ma ?nl" 
    show " ?P4 ma ?nl"     
    proof(case_tac "isDefinedIn ma nl")
      let ?P="λ ma nl enttr  . enttr ∈ nl ∧
                 ( enttr = Input ma ∨
                  (∃data .  enttr = Delay ma data ) ∨
                  (∃inps tab.  enttr = Gate ma inps tab))"
      assume d1:"isDefinedIn ma nl "

      from b3  c2  d1 have d2:"?P4 ma nl"
        apply -
        apply(drule_tac x="ma" in spec)
        by simp
      thm exE
      from d2 obtain enttr0 
        where d3: "?P ma nl enttr0  &(∀ enttr'. ?P ma nl enttr'  --> enttr' =enttr0)"
        apply -  apply(erule_tac P="?P ma nl" in  ex1E) apply blast
        done

      then have d4:"?P ma ?nl enttr0"
        by blast
      have d5:"(∀ enttr'. ?P ma ?nl enttr'  --> enttr' =enttr0)"
        apply - 
        apply(rule allI,rule impI,erule conjE)
        apply(erule UnE)
        apply(cut_tac b1 d1)
        apply(unfold isDefinedIn_def)
        apply simp
        apply(cut_tac d3,blast)
        done
      from d4 d5 show ?thesis 
        apply - 
        apply(rule_tac a="enttr0" in ex1I) 
        apply blast
        apply(drule_tac x="x" in spec)
        
        by blast
      next
        assume d1:"~isDefinedIn ma nl " 
        let ?l="(Input m)"
        from c2 have "isDefinedIn ma ({(Input m)})| isDefinedIn ma nl"
          apply - apply(unfold isDefinedIn_def)
          by blast
        with d1  have d2:"ma=m" apply - 
          apply(unfold isDefinedIn_def entityPart_def, simp)
          done
        let ?P="λ ma nl enttr  . enttr ∈ nl ∧
                 ( enttr = Input ma ∨
                  (∃data . enttr = Delay ma data ) ∨
                  (∃inps tab.  enttr = Gate ma inps tab))"
        from d2 have d4:"?P ma ?nl ?l"
          apply - apply simp
          done
        have d5:"ALL x. ?P ma ?nl x --> x=?l"
          proof(rule allI, rule impI)
            fix x
            assume e1:" ?P ma ?nl x"
            show "x=?l"
            proof -
              from e1 have e2:"x ∈ {(Input m)} |x:nl"
                by simp
              moreover
              {assume f1:"x:{(Input m)}"
                from f1 have ?thesis by simp
              }
              moreover
              {assume f1:"x:nl"
                from f1 and e1 have "isDefinedIn ma nl"
                  apply -
                  apply(unfold isDefinedIn_def)
                  apply(rule_tac x="x" in bexI)
                  apply auto
                  done
                with d1 have ?thesis by blast
              }
              ultimately
              show ?thesis by blast
            qed
          qed
          from  d2 d4 d5 show ?thesis
            apply - apply(rule_tac a="?l" in ex1I)
            apply blast
            apply(drule_tac x="x" in spec) 
            apply blast
            done
        qed
      qed
    next
      fix data m nl
      assume b1:"¬ isDefinedIn m nl" and
        b2:" nl ∈ netlists" and
        b3:"ALL m. (?P2 m nl--> ?P4 m nl)" 
      let ?nl="({(Delay m data)} ∪ nl)" 
      show "ALL ma.( ?P2 ma ?nl--> ?P4 ma ?nl)"
      proof(rule allI,(rule impI)+)
        fix ma
        assume 
          c2:"isDefinedIn ma ?nl" 
        show " ?P4 ma ?nl"     
        proof(case_tac "isDefinedIn ma nl")
          let ?P="λ ma nl enttr  . enttr ∈ nl ∧
            (enttr = Input ma ∨
            (∃data .  enttr = Delay ma data ) ∨
            (∃inps tab.  enttr = Gate ma inps tab))"
          assume d1:"isDefinedIn ma nl "

          from b3  c2  d1 have d2:"?P4 ma nl"
            apply -
            apply(drule_tac x="ma" in spec)
            by simp
          from d2 obtain enttr0 
            where d3: "?P ma nl enttr0  &(∀ enttr'. ?P ma nl enttr'  --> enttr' =enttr0)"
            apply -  apply(erule_tac P="?P ma nl" in  ex1E) apply blast
            done

          then have d4:"?P ma ?nl enttr0"
            by blast
          have d5:"(∀ enttr'. ?P ma ?nl enttr'  --> enttr' =enttr0)"
            apply - 
            apply(rule allI,rule impI,erule conjE)
            apply(erule UnE)
            apply(cut_tac b1 d1)
            apply(unfold isDefinedIn_def)
            apply simp
            apply(cut_tac d3,blast)
            done
          from d4 d5 show ?thesis 
            apply - 
            apply(rule_tac a="enttr0" in ex1I) 
            apply blast
            apply(drule_tac x="x" in spec)
            by blast
        next
          assume d1:"~isDefinedIn ma nl "       
          let ?l="(Delay m data)"
          from c2 have "isDefinedIn ma {?l}| isDefinedIn ma nl"
            apply - apply(unfold isDefinedIn_def)
            by blast
          with d1  have d2:"ma=m" apply - 
            apply(unfold isDefinedIn_def entityPart_def, simp)
            done
          let ?P="λ ma nl enttr  . enttr ∈ nl ∧
                 ( enttr = Input ma ∨
                  (∃data .  enttr = Delay ma data ) ∨
                  (∃inps tab.  enttr = Gate ma inps tab))"
          from d2 have d4:"?P ma ?nl ?l"
            apply - apply simp
            done
          have d5:"ALL x. ?P ma ?nl x --> x=?l"
          proof(rule allI, rule impI)
            fix x
            assume e1:" ?P ma ?nl x"
            show "x=?l"
            proof -
              from e1 have e2:"x ∈ {?l} |x:nl"
                by simp
              moreover
              {assume f1:"x:{?l}"
                from f1 have ?thesis by simp
              }
              moreover
              {assume f1:"x:nl"
                from f1 and e1 have "isDefinedIn ma nl"
                  apply -
                  apply(unfold isDefinedIn_def)
                  apply(rule_tac x="x" in bexI)
                  apply auto
                  done
                with d1 have ?thesis by blast
              }
              ultimately
              show ?thesis by blast
            qed
          qed
          from  d2 d4 d5 show ?thesis
            apply - apply(rule_tac a="?l" in ex1I)
            apply blast
            apply(drule_tac x="x" in spec) 
            apply blast
            done
        qed
      qed
      next
        fix inps  m nl tab
        assume b1:"¬ isDefinedIn m nl" and
          b2:" nl ∈ netlists" and
          b3:"ALL m. (?P2 m nl--> ?P4 m nl)" and
          b4:" ∀n. n mem inps --> isDefinedIn n nl "
        let ?nl="{(Gate m inps tab)} Un nl"
        show "ALL ma.( ?P2 ma ?nl--> ?P4 ma ?nl)"
        proof(rule allI,(rule impI)+)
        fix ma
        assume 
          c2:"isDefinedIn ma ?nl" 
        show " ?P4 ma ?nl"     
        proof(case_tac "isDefinedIn ma nl")
          let ?P="λ ma nl enttr  . enttr ∈ nl ∧
            ( enttr = Input ma ∨
            (∃data .  enttr = Delay ma data ) ∨
            (∃inps tab.  enttr = Gate ma inps tab))"
          assume d1:"isDefinedIn ma nl "

          from b3  c2  d1 have d2:"?P4 ma nl"
            apply -
            apply(drule_tac x="ma" in spec)
            by simp
          from d2 obtain enttr0 
            where d3: "?P ma nl enttr0  &(∀ enttr'. ?P ma nl enttr'  --> enttr' =enttr0)"
            apply -  apply(erule_tac P="?P ma nl" in  ex1E) apply blast
            done

          then have d4:"?P ma ?nl enttr0"
            by blast
          have d5:"(∀ enttr'. ?P ma ?nl enttr'  --> enttr' =enttr0)"
            apply - 
            apply(rule allI,rule impI,erule conjE)
            apply(erule UnE)
            apply(cut_tac b1 d1)
            apply(unfold isDefinedIn_def)
            apply simp
            apply(cut_tac d3,blast)
            done
          from d4 d5 show ?thesis 
            apply - 
            apply(rule_tac a="enttr0" in ex1I) 
            apply blast
            apply(drule_tac x="x" in spec)
            by blast
        next
          assume d1:"~isDefinedIn ma nl "       
          let ?l="(Gate m inps tab)"
          from c2 have "isDefinedIn ma {?l}| isDefinedIn ma nl"
            apply - apply(unfold isDefinedIn_def)
            by blast
          with d1  have d2:"ma=m" apply - 
            apply(unfold isDefinedIn_def entityPart_def, simp)
            done
          let ?P="λ ma nl enttr  . enttr ∈ nl ∧
                 ( enttr = Input ma ∨
                  (∃data .  enttr = Delay ma data ) ∨
                  (∃inps tab.  enttr = Gate ma inps tab))"
          from d2 have d4:"?P ma ?nl ?l"
            apply - apply simp
            done
          have d5:"ALL x. ?P ma ?nl x --> x=?l"
          proof(rule allI, rule impI)
            fix x
            assume e1:" ?P ma ?nl x"
            show "x=?l"
            proof -
              from e1 have e2:"x ∈ {?l} |x:nl"
                by simp
              moreover
              {assume f1:"x:{?l}"
                from f1 have ?thesis by simp
              }
              moreover
              {assume f1:"x:nl"
                from f1 and e1 have "isDefinedIn ma nl"
                  apply -
                  apply(unfold isDefinedIn_def)
                  apply(rule_tac x="x" in bexI)
                  apply auto
                  done
                with d1 have ?thesis by blast
              }
              ultimately
              show ?thesis by blast
            qed
          qed
          from  d2 d4 d5 show ?thesis
            apply - apply(rule_tac a="?l" in ex1I)
            apply blast
            apply(drule_tac x="x" in spec) 
            apply blast
            done
        qed
      qed
    qed


lemma nodeNameIsUniqueInNet:
assumes a1:"nl:netlists" and 
a2:"isDefinedIn n nl"
shows "
 ( EX! enttr.(enttr):nl 
                             &(( enttr)= Input n |
                               ( EX data .( enttr)=Delay n data )|
                               (EX inps tab.( enttr)= Gate n inps tab)))"
apply -
apply(cut_tac a1 a2)
apply(blast dest:nodeNameIsUniqueInNet)
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 InputNameDelayNameIsDiff:
assumes a1:"nl:netlists" and
  a2:"Input n:nl" and a3:"Delay n data:nl"
shows
  "False"
proof -
  from a2 have "isDefinedIn n nl"
    apply -
    apply(unfold isDefinedIn_def)
    apply(rule_tac x="Input n" in bexI)
    apply simp
    by blast
  with a1 a2 a3 show ?thesis
    apply -
    apply (drule  nodeNameIsUniqueInNet)
    apply assumption
    apply blast
    done
qed

lemma InputNameGateNameIsDiff:
assumes a1:"nl:netlists" and
  a2:"Input n:nl" and a3:"Gate n inps tab:nl"
shows
  "False"
proof -
  from a2 have "isDefinedIn n nl"
    apply -
    apply(unfold isDefinedIn_def)
    apply(rule_tac x="Input n" in bexI)
    apply simp
    by blast
  with a1 a2 a3 show ?thesis
    apply -
    apply (drule  nodeNameIsUniqueInNet)
    apply assumption
    apply blast
    done
qed

lemma DelayNameGateNameIsDiff:
assumes a1:"nl:netlists" and
  a2:"Delay n data:nl" and a3:"Gate n inps tab:nl"
shows
  "False"
proof -
  from a2 have "isDefinedIn n nl"
    apply -
    apply(unfold isDefinedIn_def)
    apply(rule_tac x="Delay n data" in bexI)
    apply simp
    by blast
  with a1 a2 a3 show ?thesis
    apply -
    apply (drule  nodeNameIsUniqueInNet)
    apply assumption
    apply blast
    done
qed

lemma DelayNameIsSame:
assumes a1:"nl:netlists" and
  a2:"Delay n data:nl" and a3:"Delay n data0:nl"
shows
  "data=data0 "
proof -
  from a2 have "isDefinedIn n nl"
    apply -
    apply(unfold isDefinedIn_def)
    apply(rule_tac x="Delay n data" in bexI)
    apply simp
    by blast
  with a1 a2 a3 show ?thesis
    apply -
    apply (drule  nodeNameIsUniqueInNet)
    apply assumption
    apply blast
    done
qed

lemma GateNameIsSame:
assumes a1:"nl:netlists" and
  a2:"Gate n inps tab:nl" and a3:"Gate n inps1 tab1:nl"
shows
  "inps=inps1& tab=tab1"
proof -
  from a2 have "isDefinedIn n nl"
    apply -
    apply(unfold isDefinedIn_def)
    apply(rule_tac x="Gate n inps tab" in bexI)
    apply simp
    by blast
  with a1 a2 a3 show ?thesis
    apply -
    apply (drule  nodeNameIsUniqueInNet)
    apply assumption
    apply blast
    done
qed


lemma stTransOfNil:" rclosure {} s={}"
apply(rule  equals0I )
apply(erule rclosure.induct)
apply auto
done




lemma stateListLeq:
shows  
  "∀stateLs1. ∀inps .  (length stateLs1 = length inps)-->
  (length stateLs = length inps)-->
  ( ∀pair0. pair0 mem zip inps stateLs -->  pair0 ∈ rclosure nl s1&
      (∀pair2. pair2 ∈ rclosure nl s2 & fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2))-->   
  ( ∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s2) -->
  stateSqLeq1 stateLs stateLs1" ( is "?P  stateLs ")

proof (induct_tac stateLs)
  show "?P []" by simp 
next
  fix a list0
  assume b1:"?P list0 "
  show "?P (a # list0) "
    proof((rule allI)+,(rule impI)+)
      fix stateLs1 inps
      assume c1:"length stateLs1 = length inps " and
        c2:" length (a # list0) = length inps" and
        c3:"  ∀pair0. pair0 mem zip inps (a # list0) --> 
        pair0 ∈ rclosure nl s1 ∧ (∀pair2. pair2 ∈ rclosure nl s2 ∧ fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2)" (is "ALL pair0. ?Ant pair0 --> ?Con1 pair0 &?Con2 pair0") and
        c4:"∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s2"
      show "stateSqLeq1 (a # list0) stateLs1"
        proof -
          
          have d1:"0<length (a # list0)" by simp
          from d1 c2 have d1:"0<length inps" 
            apply - apply arith done
          with c1 have "0<length stateLs1" by simp
          then have d2:"EX y ys. stateLs1 = y # ys"
            apply -
            apply(subgoal_tac "EX n. length  stateLs1=Suc n")
            apply(simp add:length_Suc_conv )
            by(rule arith1)
          from d1 have d3:"EX y ys. inps = y # ys"
            apply -
            apply(subgoal_tac "EX n. length  inps=Suc n")
            apply(simp add:length_Suc_conv )
            by(rule arith1)
          from d2 obtain st1 and stl1 where d4:"stateLs1 =st1 # stl1"
            by blast
          from d3 obtain inp1 and inpl1 where d5:"inps =inp1 # inpl1"
            by blast
          let ?pair0="(inp1,a)"
          let ?pair2="(inp1,st1)"
          from c4  have d6:"?pair2 :rclosure nl s2"
            apply -
            apply(drule_tac x="?pair2" in spec)
            apply(cut_tac d4 d5)
            apply simp
            done
          from c3 have d7:"( ?Con1 ?pair0 & ?Con2 ?pair0 )"
            apply -
            apply(drule_tac x="?pair0" in spec)
            apply(cut_tac d4 d5)
            apply auto
            done
          show ?thesis
          proof( simp only:stateSqLeq1Cons,rule conjI)
            from d7 d4 show "a \<sqsubseteq>  hd stateLs1"
              apply -
              apply(erule conjE)
              apply(drule_tac x="?pair2" in spec)
              apply(cut_tac d6,simp)
              done
          next
            have f1:"length stl1 = length inpl1"
              apply(cut_tac c1 c2 d4 d5,simp)
              done
            have f2:"length list0 = length inpl1"
              by(cut_tac c1 c2 d4 d5,simp)

            have f3:" ∀pair0. pair0 mem zip inpl1 ( list0) -->  pair0 ∈ rclosure nl s1&
              (∀pair2. pair2 ∈ rclosure nl s2 & fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2)"
              proof(rule allI,rule impI)
                fix pair0
                assume e1:"pair0 mem zip inpl1 list0"
                show "pair0 ∈ rclosure nl s1 ∧ (∀pair2. pair2 ∈ rclosure nl s2 ∧ fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2)"
                  proof -
                    have "pair0 mem zip inps (a#list0) "
                      apply -
                      apply(cut_tac c1 c2 d4 d5 e1)
                      apply simp
                      done
                    then show ?thesis
                      apply -
                      apply(cut_tac c3)
                      apply(drule_tac x="pair0" in spec)
                      apply blast
                      done
                  qed
                qed

            have f4:" (∀pair0. pair0 mem zip inpl1 stl1 --> pair0 ∈ rclosure nl s2)"
            proof(rule allI,rule impI)
              fix pair0
              assume e1:"pair0 mem zip inpl1 stl1"
              show" pair0 ∈ rclosure nl s2"
              proof -
                have "pair0 mem zip inps stateLs1 "
                  apply -
                  apply(cut_tac c1 c2 d4 d5 e1)
                  apply simp
                  done
                then show ?thesis
                  apply -
                  apply(cut_tac c4)
                  apply(drule_tac x="pair0" in spec)
                  apply blast
                  done
              qed
            qed
            show "stateSqLeq1 list0 (tl stateLs1)"
              apply(cut_tac b1) 
              apply(drule_tac x="stl1" in spec)
              apply(drule_tac x="inpl1" in spec)
              apply(cut_tac f1 f2 f3 f4 d4 d5)
              apply auto  
              done
          qed
        qed
      qed
    qed

lemma stateListLeq:
  assumes
  a1:"(length stateLs1 = length inps)" and
  a2:"(length stateLs = length inps)" and
  a3:"
  ( ∀pair0. pair0 mem zip inps stateLs -->  pair0 ∈ rclosure nl s1&
      (∀pair2. pair2 ∈ rclosure nl s2 & fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2))" and
  a4:"
  ( ∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s2) "
  shows
 " stateSqLeq1 stateLs stateLs1"
  apply(cut_tac stateListLeq[where stateLs="stateLs" 
    and ?s1.0="s1" and ?s2.0="s2" and nl="nl"])
  apply(cut_tac a1 a2 a3 a4,blast)
done

lemma ftabIsMono:
assumes 
  a1:"∀pair0. pair0 mem zip inps stateLs -->
  pair0 ∈ rclosure nl s1 ∧
  (∀pair2. pair2 ∈ rclosure nl s2 --> fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2)" and
  a2:" length stateLs = length inps" and
  a3:" ∀pair0. pair0 mem zip inps stateLsa -->
  pair0 ∈ rclosure nl s2" and
  a4:"length stateLsa = length inps"

shows " (∀ l. (l∈ (set tab)) --> length l=length inps)-->
  funOfTab tab stateLs \<sqsubseteq> funOfTab tab stateLsa" (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(rule impI)
    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 ftabIsMono:
assumes 
  a1:"∀pair0. pair0 mem zip inps stateLs -->
  pair0 ∈ rclosure nl s1 ∧
  (∀pair2. pair2 ∈ rclosure nl s2 --> fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2)" and
  a2:" length stateLs = length inps" and
  a3:" ∀pair0. pair0 mem zip inps stateLsa -->
  pair0 ∈ rclosure nl s2" and
  a4:"length stateLsa = length inps" and
  a5:"(∀ l. (l∈ (set tab)) --> length l=length inps)"

shows " funOfTab tab stateLs \<sqsubseteq> funOfTab tab stateLsa"
by(cut_tac a1 a2 a3 a4 a5 ftabIsMono,auto)




lemma rclosureMono:
assumes 
  a1:"nl:netlists"
shows
  "  ALL s1 s2. s1  \<sqsubseteq>s s2 -->
  ( pair1:rclosure nl s1 -->(ALL pair2. pair2:rclosure nl s2 --> (fst pair1)=fst pair2
 -->  (snd pair1 )\<sqsubseteq> snd pair2))"
apply -
apply(rule allI)+
apply(rule impI)
apply(rule impI)
apply(erule rclosure.induct)
apply(rule allI)
apply(rule impI)
apply(erule rclosure.induct)
apply(unfold stateLeq_def)
apply simp
apply(rule impI)
apply(subgoal_tac "x=n")
apply simp 
apply simp
apply(rule impI)
apply(subgoal_tac "x=n")
apply simp 
apply(cut_tac a1)
apply(blast dest: InputNameGateNameIsDiff)
apply simp

apply(rule allI)
apply(rule impI)
apply(erule rclosure.induct)
apply simp
apply(rule impI)
apply(subgoal_tac "n=na")
apply simp
apply simp
apply(rule impI)
apply(subgoal_tac "n=na")
apply simp
apply(cut_tac a1)
apply(blast dest: DelayNameGateNameIsDiff)
apply simp

apply(rule allI)
apply(rule impI)
apply(erule rclosure.induct)
apply(rule impI)
apply(subgoal_tac "n=x")
apply simp
apply(cut_tac a1,blast dest: InputNameGateNameIsDiff)
apply simp
apply(rule impI)
apply(subgoal_tac "n=na")
apply simp
apply(cut_tac a1)
apply(blast dest: DelayNameGateNameIsDiff)
apply simp
apply(rule impI)
apply(subgoal_tac "n=na")
prefer 2
apply simp
apply(subgoal_tac "inps=inpsa & tab=taba")
prefer 2
apply(cut_tac a1,rule GateNameIsSame)
apply simp
apply simp
apply simp
apply(simp (no_asm))
apply(rule lubMono)
apply(subgoal_tac "funOfTab tab stateLs \<sqsubseteq>  funOfTab tab stateLsa")
apply simp
apply(rule ftabIsMono)
apply assumption+
apply blast
apply simp
apply simp
apply simp
done

lemma rclosureMonoAny:
assumes 
  a1:"nl:netlists"
shows
  "  ALL s1 s2. (ALL n. isDefinedIn n nl -->  ((s1 n ) \<sqsubseteq> ( s2 n)))
        -->( pair1:rclosure nl s1 -->(ALL pair2. pair2:rclosure nl s2 --> (fst pair1)=fst pair2
        -->  (snd pair1 )\<sqsubseteq> snd pair2))"
apply -
apply(rule allI)+
apply(rule impI)
apply(rule impI)
apply(erule rclosure.induct)
apply(rule allI)
apply(rule impI)
apply(erule rclosure.induct)
apply simp
apply(rule impI)
apply(subgoal_tac "isDefinedIn xa nl")
apply simp 
apply(unfold isDefinedIn_def,simp)
apply(rule_tac x="Input xa " in bexI)
apply( simp,simp)
apply simp
apply(rule impI,simp)
apply(cut_tac a1)
apply(blast dest: InputNameDelayNameIsDiff)
apply simp
apply(rule impI,simp)
apply(cut_tac a1)
apply(blast dest: InputNameGateNameIsDiff)

apply(fold isDefinedIn_def)
apply(rule allI)
apply(rule impI)
apply(erule rclosure.induct)
apply simp
apply(rule impI)
apply simp
apply(cut_tac a1)
apply(blast dest: InputNameDelayNameIsDiff)
apply simp
apply(rule impI)
apply(subgoal_tac "isDefinedIn na nl")
apply simp 
apply(unfold isDefinedIn_def,simp)
apply(rule_tac x="Delay na dataa " in bexI)
apply( simp,simp)
apply(rule impI,simp)
apply(cut_tac a1)
apply(blast dest: DelayNameGateNameIsDiff)

apply(fold isDefinedIn_def)
apply(rule allI)
apply(rule impI)
apply(erule rclosure.induct)
apply(rule impI)
apply(subgoal_tac "n=x")
apply simp
apply(cut_tac a1,blast dest: InputNameGateNameIsDiff)
apply simp
apply(rule impI)
apply(subgoal_tac "n=na")
apply simp
apply(cut_tac a1)
apply(blast dest: DelayNameGateNameIsDiff)
apply simp
apply(rule impI)
apply(subgoal_tac "n=na")
prefer 2
apply simp
apply(subgoal_tac "inps=inpsa & tab=taba")
prefer 2
apply(cut_tac a1,rule GateNameIsSame)
apply simp
apply simp
apply simp
apply(simp (no_asm))
apply(rule lubMono)
apply(subgoal_tac "funOfTab tab stateLs \<sqsubseteq>  funOfTab tab stateLsa")
apply simp
apply(rule ftabIsMono)
apply assumption+
apply blast
apply simp
apply simp
apply(subgoal_tac "isDefinedIn na nl")
apply simp 
apply(unfold isDefinedIn_def,simp)
apply(rule_tac x="Gate na inpsa taba  " in bexI)
apply( simp,simp)
done

lemma rclosureMono:
assumes 
  a1:"nl:netlists" and
  a2:"s1  \<sqsubseteq>s s2" and
  a3:" pair1:rclosure nl s1" and
  a4:" pair2:rclosure nl s2 " and 
  a5:"(fst pair1)=fst pair2" 
shows
  "   (snd pair1 )\<sqsubseteq> snd pair2"
apply(cut_tac  rclosureMono[where nl="nl" and ?pair1.0="pair1" ])
apply(cut_tac a1 a2 a3 a4 a5)
apply(drule_tac x="s1" in spec)
apply(drule_tac x="s2" in spec)
apply blast
apply assumption
done

lemma rclosureMonoAny:
assumes 
  a1:"nl:netlists" and
  a2:"(ALL n. isDefinedIn n nl -->  ((s1 n ) \<sqsubseteq> ( s2 n)))" and
  a3:" pair1:rclosure nl s1" and
  a4:" pair2:rclosure nl s2 " and 
  a5:"(fst pair1)=fst pair2" 
shows
  "   (snd pair1 )\<sqsubseteq> snd pair2"
apply(cut_tac  rclosureMonoAny[where nl="nl" and ?pair1.0="pair1" ])
apply(cut_tac a1 a2 a3 a4 a5)
apply(drule_tac x="s1" in spec)
apply(drule_tac x="s2" in spec)
apply blast
apply assumption
done
(*
lemma rclosureAdd:
  assumes
  a1:"pair:rclosure ({l}Un nl) s" and
  a2:"fanout ( l)=n" and 
  a3:"~isDefinedIn n nl"and
  a4:"nl:netlists"
  shows
  "pair:rclosure  nl s | fst pair =n" (is "?P pair")
  using a1
proof(induct)
  fix x
  assume b1:"Input x ∈ {l} ∪ nl "
  show "?P (x, s x)"
    by(cut_tac b1 a2, auto dest:stAddInput)
next
  fix data na
  assume b1:"  Delay na data∈ {l} ∪ nl "
  show "?P (na, s na)"
    by(cut_tac b1 a2, auto dest:stAddDelay)
next
  fix inps na stateLs tab
  assume b1:"Gate na inps tab ∈ {l} ∪ 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 ({l} ∪ nl) s ∧ (pair0 ∈ rclosure nl s ∨ fst pair0 = n)"
  show "?P (na, lub (funOfTab tab stateLs) (s na))"
    proof -
      from b1 have c1:"l=Gate na inps tab |Gate na inps tab:nl "
        by blast
      moreover
      {assume c1:"l=Gate na inps tab"
        from c1 a2 have "n=na" by auto   
        then   have ?thesis  
          by simp
      }
      moreover
      {assume c1:"Gate na inps tab:nl "
        from c1 
        have c2:"ALL pair0. pair0 mem zip inps stateLs -->isDefinedIn (fst pair0) nl" 
          apply -
          apply(rule allI,rule impI)
          apply(cut_tac b4)
          apply(subgoal_tac "(pair0 ∈ rclosure nl s ∨ fst pair0 = n)")
          apply(erule disjE)
        have ?thesis
          apply(rule disjI1)
          apply(rule stAddGate)
          apply assumption+
          apply(rule allI,rule impI)
          apply(cut_tac b4)
          apply(drule_tac x="pair0" in spec)
          apply(drule mp)
          apply(simp)
          apply(erule conjE, erule disjE)
          apply simp      
          apply(cut_tac c2 a3)
          apply blast
          done
      }
      ultimately show ?thesis by blast
    qed
  qed*)



lemma inputNodeIsUniInFclosure :
  assumes a1:"pair0 ∈ rclosure ( nl) s" and
  
  a3:"(Input n ):nl" and
  a4:"nl:netlists"
  shows  " (fst pair0 )= n --> pair0 = (n, s n)"
  using a1
  proof(induct)
    fix x
    
    show "fst (x, s x) = n --> (x, s x) = (n, s n)"
      apply -
      apply auto
      done
    next
      fix data na
      show "fst (na, s na) = n --> (na, s na) = (n, s n)"
        by auto
    next
      fix inps na stateLs tab 
      assume b1:"Gate na inps tab ∈  nl"
     
      show " fst (na, lub (funOfTab tab stateLs) (s na)) = n --> 
        (na, lub (funOfTab tab stateLs) (s na)) = (n, s n)"
        apply(rule impI)
        apply(cut_tac a3 b1 a4)
        apply(auto) thm nodeNameIsUniqueInNet
        apply(subgoal_tac "isDefinedIn n nl")
        apply(blast dest: nodeNameIsUniqueInNet)
        apply(unfold isDefinedIn_def)
        apply(rule_tac x="Input n" in bexI)
        apply simp
        apply assumption
        done
    qed

lemma latchNodeIsUniInFclosure :
  assumes a1:"pair0 ∈ rclosure  nl s" and
  a3:"(Delay  n data):nl" and
  a4:"nl:netlists"
  shows  " (fst pair0 )= n --> pair0 = (n, s n)"
  using a1
  proof(induct)
    fix x
    show "fst (x, s x) = n --> (x, s x) = (n, s n)"
      apply -
      apply auto
      done
    next
      fix data na
      show "fst (na, s na) = n --> (na, s na) = (n, s n)"
        by auto
    next
      fix inps na stateLs tab 
      assume b1:"Gate na inps tab ∈  nl"

      show " fst (na, lub (funOfTab tab stateLs) (s na)) = n --> 
        (na, lub (funOfTab tab stateLs) (s na)) = (n, s n)"
        apply(rule impI)
        apply(cut_tac a3 b1 a4)
        apply(auto) thm nodeNameIsUniqueInNet
        apply(subgoal_tac "isDefinedIn n nl")
        apply(blast dest: nodeNameIsUniqueInNet)
        apply(unfold isDefinedIn_def)
        apply(rule_tac x="(Delay  n data)" in bexI)
        apply simp
        apply assumption
        done
    qed

lemma stateListEq:
shows  
  "∀stateLs1. ∀inps .  (length stateLs1 = length inps)-->
  (length stateLs = length inps)-->
  (∀pair0. pair0 mem zip inps stateLs --> 
           pair0 ∈ rclosure nl s ∧ (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0))-->   
  ( ∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s) -->
  stateLs=stateLs1" ( is "?P  stateLs ")
proof (induct_tac stateLs)
  show "?P []" by simp 
next
  fix a list0
  assume b1:"?P list0 "
  show "?P (a # list0) "
    proof((rule allI)+,(rule impI)+)
      fix stateLs1 inps
      assume c1:"length stateLs1 = length inps " and
        c2:" length (a # list0) = length inps" and
        c3:"  ∀pair0. pair0 mem zip inps (a # list0) --> 
        pair0 ∈ rclosure nl s ∧
        (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0 )"
        (is "ALL pair0. ?Ant pair0 --> ?Con1 pair0 &?Con2 pair0") and
        c4:"∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s"
      show " (a # list0)= stateLs1"
        proof -
          
          have d1:"0<length (a # list0)" by simp
          from d1 c2 have d1:"0<length inps" 
            apply - apply arith done
          with c1 have "0<length stateLs1" by simp
          then have d2:"EX y ys. stateLs1 = y # ys"
            apply -
            apply(subgoal_tac "EX n. length  stateLs1=Suc n")
            apply(simp add:length_Suc_conv )
            by(rule arith1)
          from d1 have d3:"EX y ys. inps = y # ys"
            apply -
            apply(subgoal_tac "EX n. length  inps=Suc n")
            apply(simp add:length_Suc_conv )
            by(rule arith1)
          from d2 obtain st1 and stl1 where d4:"stateLs1 =st1 # stl1"
            by blast
          from d3 obtain inp1 and inpl1 where d5:"inps =inp1 # inpl1"
            by blast
          let ?pair0="(inp1,a)"
          let ?pair2="(inp1,st1)"
          from c4  have d6:"?pair2 :rclosure nl s"
            apply -
            apply(drule_tac x="?pair2" in spec)
            apply(cut_tac d4 d5)
            apply simp
            done
          from c3 have d7:"( ?Con1 ?pair0 & ?Con2 ?pair0 )"
            apply -
            apply(drule_tac x="?pair0" in spec)
            apply(cut_tac d4 d5)
            apply auto
            done
          let ?R="λ pair. pair ∈ rclosure nl s ∧ fst pair = inp1"
          show ?thesis
          proof( cut_tac d4,simp,rule conjI)
            from d7 d4 show "a =st1"
              apply -
              apply(erule conjE)
              apply simp
              apply(subgoal_tac "?R ?pair0")
              prefer 2 
              apply simp
              apply(subgoal_tac "?R ?pair2")
              prefer 2 
              apply(cut_tac d6, simp)
              apply blast
              done
          next
            have f1:"length stl1 = length inpl1"
              apply(cut_tac c1 c2 d4 d5,simp)
              done
            have f2:"length list0 = length inpl1"
              by(cut_tac c1 c2 d4 d5,simp)

            have f3:" ∀pair0. pair0 mem zip inpl1 ( list0) -->  pair0 ∈ rclosure nl s&
              (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0 )" 
              (is "ALL pair0. ?Ant pair0 --> ?Con1 pair0 &?Con2 pair0")
              proof(rule allI,rule impI)
                fix pair0
                assume e1:"pair0 mem zip inpl1 list0"
                show "?Con1 pair0 &?Con2 pair0"
                  proof -
                    have "pair0 mem zip inps (a#list0) "
                      apply -
                      apply(cut_tac c1 c2 d4 d5 e1)
                      apply simp
                      done
                    then show ?thesis
                      apply -
                      apply(cut_tac c3)
                      apply(drule_tac x="pair0" in spec)
                      apply simp
                      
                      done
                  qed
                qed

            have f4:" (∀pair0. pair0 mem zip inpl1 stl1 --> pair0 ∈ rclosure nl s)"
            proof(rule allI,rule impI)
              fix pair0
              assume e1:"pair0 mem zip inpl1 stl1"
              show" pair0 ∈ rclosure nl s"
              proof -
                have "pair0 mem zip inps stateLs1 "
                  apply -
                  apply(cut_tac c1 c2 d4 d5 e1)
                  apply simp
                  done
                then show ?thesis
                  apply -
                  apply(cut_tac c4)
                  apply(drule_tac x="pair0" in spec)
                  apply blast
                  done
              qed
            qed
            show "list0 = stl1"
              apply(cut_tac b1) 
              apply(drule_tac x="stl1" in spec)
              apply(drule_tac x="inpl1" in spec)
              apply(cut_tac f1 f2 f3 f4 d4 d5)
              apply auto  
              done
          qed
        qed
      qed
    qed

lemma stateListEq:
  assumes
  a1:"(length stateLs1 = length inps)" and
  a2:"(length stateLs = length inps)" and
  a3:"(∀pair0. pair0 mem zip inps stateLs --> 
           pair0 ∈ rclosure nl s ∧ (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0))" and
  a4:"
  ( ∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s) "
  shows
 "  stateLs = stateLs1"
  apply(cut_tac stateListEq[where stateLs="stateLs"])
  apply(cut_tac a1 a2 a3 a4,blast)
done

lemma GateNodeIsUniInFclosure :
  assumes a1:"pair0 ∈ rclosure  nl s" and 
  a2:"length stateLs=length inps" and
  a3:"(Gate  n inps tab):nl" and
  a4:"nl:netlists"  and
  a5:"(∀pair0. pair0 mem zip inps stateLs --> 
           pair0 ∈ rclosure nl s ∧ (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0))"
  shows  " (fst pair0 )= n --> pair0 = (n,  lub (funOfTab tab stateLs) (s n))" 
  using a1
  proof(induct)
    fix x
    assume b1:"Input x ∈ nl"
    show "fst (x, s x) = n --> (x, s x) = (n,  lub (funOfTab tab stateLs) (s n))"
      apply -
      apply(rule impI)
      apply(cut_tac a3 b1 a4)
      apply(auto) thm nodeNameIsUniqueInNet
      apply(subgoal_tac "isDefinedIn n nl")
      apply(blast dest: nodeNameIsUniqueInNet)
      apply(unfold isDefinedIn_def)
      apply(rule_tac x="(Gate  n inps tab)" in bexI)
      apply simp
      apply assumption      
      done
    next
      fix data na
      assume b1:"Delay na data ∈ nl"
      show "fst (na, s na) = n --> (na, s na) = (n,  lub (funOfTab tab stateLs) (s n))"
        apply -
        apply(rule impI)
        apply(cut_tac a3 b1 a4)
        apply(auto) thm nodeNameIsUniqueInNet
        apply(subgoal_tac "isDefinedIn n nl")
        apply(blast dest: nodeNameIsUniqueInNet)
        apply(unfold isDefinedIn_def)
        apply(rule_tac x="(Gate  n inps tab)" in bexI)
        apply simp
        apply assumption      
        done
    next
      fix inpsa na stateLsa taba 
      assume b1:"Gate na inpsa taba ∈  nl" and 
        b2:"length stateLsa = length inpsa" and
        b3:"∀l. l ∈ set taba --> length l = length inpsa" and
        b4:"∀pair0. pair0 mem zip inpsa stateLsa -->
                pair0 ∈ rclosure nl s ∧ (fst pair0 = n --> pair0 = (n, lub (funOfTab tab stateLs) (s n)))"
      show "fst (na, lub (funOfTab taba stateLsa) (s na)) = n -->
          (na, lub (funOfTab taba stateLsa) (s na)) = (n, lub (funOfTab tab stateLs) (s n))"
      proof
        assume c1:"fst (na, lub (funOfTab taba stateLsa) (s na)) = n"

        from c1 have c0:"na=n" by simp
        
        with b1 a3 a4 
        have c2:"inpsa=inps & taba=tab"
         apply - 
         apply(rule GateNameIsSame)
         apply assumption+
         apply simp
         done

       have "stateLs=stateLsa" thm stateListEq
         apply(cut_tac c2 a2 b2 b4 a5)
         apply(rule_tac inps="inpsa" and nl="nl" in stateListEq)         
         apply simp
         apply simp
         apply auto
         done
        

      then show "
          (na, lub (funOfTab taba stateLsa) (s na)) = (n, lub (funOfTab tab stateLs) (s n))"
        
        
        apply(cut_tac a3 b1 a4 c0 c2)
        apply(auto) 
        done
    qed 
  qed

lemma GateNodeIsUniInFclosure :
  assumes a1:"pair0 ∈ rclosure  nl s" and 
  a2:"length stateLs=length inps" and
  a3:"(Gate  n inps tab):nl" and
  a4:"nl:netlists"  and
  a5:"(∀pair0. pair0 mem zip inps stateLs --> 
           pair0 ∈ rclosure nl s ∧ (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0))" and
  a6:"(fst pair0 )= n "
  shows  "  pair0 = (n,  lub (funOfTab tab stateLs) (s n))"
  
  thm GateNodeIsUniInFclosure
  apply(cut_tac GateNodeIsUniInFclosure[where ?pair0.0=" pair0" and n="n"
    and inps="inps" and stateLs="stateLs" and s="s"])
  apply(cut_tac a6)
  apply blast
  apply assumption+
  done


lemma pairIsUniInrclosure:
  assumes 
  a1:"nl:netlists" and 
  a2:" pair0:rclosure nl s "  
  shows  "EX! pair. pair:rclosure nl s & (fst pair)=fst pair0 " (is "EX! pair. ?P pair pair0")
  using a2
  proof(induct)
    fix x
    assume b1:" Input x ∈ nl"
    show "EX! pair. ?P pair (x, s x)"
      proof -
        let ?pair="(x, s x)"
        have c1:"?P ?pair (x, s x)"
          apply(rule conjI)
          apply(rule stAddInput)
          apply assumption
          by auto
        have "∀ pair0. ?P pair0 (x, s x) --> pair0=?pair"
          apply(rule allI,rule impI)
          apply(erule conjE) thm impI    
          apply(cut_tac b1 a1)
          apply simp
          apply(drule inputNodeIsUniInFclosure )
          apply blast+
          done
        with c1 show ?thesis
          by blast
      qed
    prefer 2
    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 s ∧ (EX! pair. ?P pair pair0)"
      show "EX! pair. ?P pair  (n, lub (funOfTab tab stateLs) (s n))"
        proof -
          let ?pair="(n, lub (funOfTab tab stateLs) (s n))"
          have c1:"?P ?pair ?pair"
            apply(rule conjI)
            apply(rule stAddGate)
            apply assumption+
            apply(cut_tac b4)
            apply blast
            by auto
          have "∀ pair0. ?P pair0 ?pair --> pair0=?pair"
            apply(rule allI,rule impI)
            apply(erule conjE) thm impI  
            apply(rule GateNodeIsUniInFclosure)
            apply assumption+
            apply simp      
            done
          with c1 show ?thesis
            by blast
      qed
    next
      
      fix data x
      assume b1:" Delay x data  ∈ nl"
      show "EX! pair. ?P pair (x, s x)"
      proof -
        let ?pair="(x, s x)"
        have c1:"?P ?pair (x, s x)"
          apply(rule conjI)
          apply(rule stAddDelay)
          apply assumption
          by auto
        have "∀ pair0. ?P pair0 (x, s x) --> pair0=?pair"
          apply(rule allI,rule impI)
          apply(erule conjE) thm impI    
          apply(cut_tac b1 a1)
          apply simp  thm latchNodeIsUniInFclosure 
          apply(drule latchNodeIsUniInFclosure )
          apply blast+
          done
        with c1 show ?thesis
          by blast
      qed       
    qed

lemma stClosureMono:
  "[|x:rclosure nl s; nl ⊆   nl0 |]==> x:rclosure nl0 s"
apply(erule rclosure.induct)
apply (auto intro:stAddInput stAddDelay stAddGate)
done

 lemma stClosureMono1:
  "[| nl ⊆   nl0 |]==> rclosure nl s ⊆ rclosure nl0 s"
apply(rule subsetI)
apply(rule stClosureMono)
apply assumption+
done

lemma aux1:
  "(ALL n. n mem inps --> (EX pair.  pair ∈ rclosure nl s ∧ fst pair = n))-->
  pair0 mem zip inps (map (λn. snd (SOME pair. pair ∈ rclosure nl s ∧ fst pair = n)) inps)-->
  pair0 ∈ rclosure ( nl) s" (is "?P inps")
  proof(induct_tac inps)
    show "?P []" by simp
  next
    fix a list
    assume b1:"?P list"
    show "?P (a#list)"
      proof((rule impI)+)
        assume 
          c1:"∀n. n mem a # list --> (∃pair. pair ∈ rclosure nl s ∧ fst pair = n)" and
          c2:"pair0 mem
          zip (a # list) (map (λn. snd (SOME pair. pair ∈ rclosure nl s ∧ fst pair = n)) (a # list))"
        show "pair0 ∈ rclosure nl s"
          proof -
            let ?P="λ  pair. pair ∈ rclosure nl s ∧ fst pair = n"
            
           
            have d1:"(EX v. (a,v):rclosure nl s)"
              apply(cut_tac c1,simp)
              done
            have d2:"ALL n. n mem list --> (∃pair. pair ∈ rclosure nl s ∧ fst pair = n)"
               apply(cut_tac c1,simp)
              done
            from c2 show ?thesis 
              
              apply -
              apply simp
              apply(case_tac "(a, snd (SOME pair. pair ∈ rclosure nl s ∧ fst pair = a)) = pair0")
              apply simp
              apply(subgoal_tac "(a, snd (SOME pair. pair ∈ rclosure nl s ∧ fst pair = a)):
                rclosure nl s")
              apply simp thm someI2_ex
              apply(rule  someI2_ex)
              apply(cut_tac d1,auto)
              apply(cut_tac c2 d2 b1)
              apply auto
              done
          qed
        qed
      qed
  
lemma aux1:
  "[|(ALL n. n mem inps --> (EX pair.  pair ∈ rclosure nl s ∧ fst pair = n));
  pair0 mem zip inps (map (λn. snd (SOME pair. pair ∈ rclosure nl s ∧ fst pair = n)) inps)|]==>
  pair0 ∈ rclosure ( nl) s"
  apply(cut_tac aux1[where inps="inps"])
  apply auto
done

lemma nodeNameisUniInrclosure1:
  assumes 
  a1:"nl:netlists" 
  shows  "ALL n. isDefinedIn n nl-->(EX pair. pair:rclosure nl s & (fst pair)=n) "
  (is "ALL n. ?Q n nl-->(EX pair. ?P pair nl n)")  
  using a1
  proof(induct)
    show "ALL n.?Q n {}-->(EX pair. ?P pair {} n)"
      by(unfold isDefinedIn_def,auto)
  next
    fix na nl
    assume b1:"¬ isDefinedIn na nl" and
      b2:"nl ∈ netlists" and
      b3:"ALL n. ?Q n nl-->(EX pair. ?P pair nl n)"  
    let ?l="Input na"
    let ?nl="{?l} Un nl"
    show "ALL n. ?Q n ?nl-->(EX pair. ?P pair ?nl n)"
      proof(rule allI,rule impI)
        fix n
        assume c1:" isDefinedIn n ({Input na} ∪ nl) "
        show "(EX pair. ?P pair ?nl n)"
        proof -
          from c1
          have "isDefinedIn n nl | na=n"
            apply -
            apply(unfold isDefinedIn_def,auto)
            done
          moreover
          {assume c1:"isDefinedIn n nl"
            from c1 b3 have "EX pair. ?P pair nl n"
              by blast
            then obtain pair where c2:"?P pair nl n"
              by blast
            then have c3:"?P pair ?nl n"
              apply(auto dest:stClosureMono)
              done
            then have ?thesis by blast
          }
          moreover
          {assume c1: " na=n"
            from c1 have "?P (na, s na) ?nl n"
              apply -
              apply(rule conjI)
              apply(blast intro:stAddInput)
              apply simp
              done
            then have ?thesis by blast
          }
          ultimately  show ?thesis by blast
        qed
      qed
    next
      fix data na nl
      assume b1:"¬ isDefinedIn na nl" and
        b2:"nl ∈ netlists" and
        b3:"ALL n. ?Q n  nl-->(EX pair. ?P pair nl n)"  
      let ?l="Delay na data"
      let ?nl="{?l} Un nl"
      show "ALL n. ?Q n ?nl-->(EX pair. ?P pair ?nl n)"
      proof(rule allI, rule impI)
        fix n
        assume c1:" isDefinedIn n ?nl "
        show "(EX pair. ?P pair ?nl n)"
        proof -
          from c1
          have "isDefinedIn n nl | na=n"
            apply -
            apply(unfold isDefinedIn_def,auto)
            done
          moreover
          {assume c1:"isDefinedIn n nl"
            from c1 b3 have "EX pair. ?P pair nl n"
              by blast
            then obtain pair where c2:"?P pair nl n"
              by blast
            then have c3:"?P pair ?nl n" 
              apply(auto dest:stClosureMono)
              done 
            then have ?thesis by blast
          }
          moreover
          {assume c1: " na=n"
            from c1 have "?P (na, s na) ?nl n"
              apply -
              apply(rule conjI)
              apply(blast intro:stAddDelay)
              apply simp
              done
            then have ?thesis by blast
          }
          ultimately  show ?thesis by blast
        qed
      qed
    next
      fix inps na nl tab
      assume b1:"¬ isDefinedIn na nl" and
        b2:"nl ∈ netlists" and
        b3:"ALL n. ?Q n nl-->(EX pair. ?P pair nl n)"  and
        b4:"∀l. l ∈ set (tab :: LIT list list) --> length l = length inps" and
        b5:"∀n. n mem inps --> isDefinedIn n nl"
      let ?l="Gate na inps tab"
      let ?nl="{?l} Un nl"
      show "ALL n. ?Q n ?nl-->(EX pair. ?P pair ?nl n)"
      proof(rule allI,rule impI)
        fix n
        assume c1:" isDefinedIn n ?nl "
        show "(EX pair. ?P pair ?nl n)"
        proof -
          from c1
          have "isDefinedIn n nl | na=n"
            apply -
            apply(unfold isDefinedIn_def,auto)
            done
          moreover
          {assume c1:"isDefinedIn n nl"
            from c1 b3 have "EX pair. ?P pair nl n"
              by blast
            then obtain pair where c2:"?P pair nl n"
              by blast
            then have c3:"?P pair ?nl n"
              apply(auto dest:stClosureMono)
              done 
            then have ?thesis by blast
          }
          moreover
          {assume c1: " na=n"
            let ?stateLs="map (λ n. snd (SOME  pair. (?P pair nl n))) inps"
            from b3 b5
            have c2:"ALL n. n mem inps --> (EX pair. ?P pair nl n)"
              apply -
              apply(rule allI,rule impI)
              by blast
            from c1 b4 b5 have "?P (na,  lub (funOfTab tab ?stateLs) (s na) ) ?nl n"
              apply -
              apply(rule conjI) thm stAddGate
              apply(rule_tac inps="inps" and stateLs="?stateLs" in  stAddGate)
              apply simp
              apply(simp add:length_map)
              apply assumption
              apply(rule allI,rule impI)
              apply(subgoal_tac "pair0 ∈ rclosure  nl s")
              apply(blast intro:stClosureMono) thm aux1
              apply(rule_tac inps="inps" in  aux1)
              apply(cut_tac c2,assumption+)
              apply simp
              done
            then have ?thesis by blast
          }
          ultimately  show ?thesis by blast
        qed
      qed
    qed
                     

lemma nodeNameisUniInrclosure:
  assumes 
  a1:"nl:netlists" and a2:" isDefinedIn n nl"
  shows  "(EX! pair. pair:rclosure nl s & (fst pair)=n) "
  (is "(EX! pair. ?P pair nl n)")
proof -
  have b1:"  EX pair. ?P pair nl n"
    apply(cut_tac nodeNameisUniInrclosure1[where nl="nl" and s="s"])
    apply(cut_tac a2, blast )
    by assumption

  then obtain pair0 where b1:"?P pair0 nl n"
    by blast
  have b2:"ALL pair. ?P pair nl n --> pair=pair0"
    apply(rule allI,rule impI)
    thm pairIsUniInrclosure
    apply(cut_tac a1 b1)
    apply(blast dest: pairIsUniInrclosure)
    done
  
  
    
  with b1 show ?thesis by blast
qed
  
  
              
            
  
  
lemma curFunMono:
assumes a1:"nl:netlists"
shows "ALL s1 s2 n. s1  \<sqsubseteq>s s2 -->( isDefinedIn n nl)--> fclosure nl s1 n \<sqsubseteq>fclosure nl s2 n"
  (is "ALL s1 s2 n. ?P s1 s2 n nl")
proof((rule allI)+,(rule impI)+)
  fix s1 s2 n
  assume b1:"isDefinedIn n nl" and
  a2:"s1  \<sqsubseteq>s s2"
  show "fclosure nl s1 n \<sqsubseteq>  fclosure nl s2 n"
    proof -
      from a1 b1 
      have b2:"( EX! pair. pair:rclosure nl s1 & (fst pair)=n )" (is "EX! enttr. ?P1 enttr ")
        apply -
        apply(drule nodeNameisUniInrclosure )
        apply blast+
        done
      from this obtain enttr1 
        where b3: "?P1  enttr1  &(∀ enttr'. ?P1  enttr'  --> enttr' =enttr1)"
        apply -  apply(erule_tac P="?P1" in  ex1E) apply blast
        done
      from a1 b1 
      have b4:"( EX! pair. pair:rclosure nl s2 & (fst pair)=n )" (is "EX! enttr. ?P2 enttr ")
        apply -
        apply(drule nodeNameisUniInrclosure )
        apply blast+
        done
      from this obtain enttr2 
        where b5: "?P2  enttr2  &(∀ enttr'. ?P2  enttr'  --> enttr' =enttr2)"
        apply -  apply(erule_tac P="?P2" in  ex1E) apply blast
        done

      
     (* let et=THE pair0. (pair0):  (rclosure enttrs  s)
                             &(fst pair0) =n in
                            snd  ( et)
      let ?Q=λ  pair0. (pair0):  (rclosure enttrs  s)
                             &(fst pair0) =n in
                            snd  ( et)
      let ?Q="λenttr. (let et = enttr
        in transPart et s1) \<sqsubseteq> 
        (let et =enttr
        in transPart et s2)"*)
      let ?Q="λ enntr. (snd enntr) \<sqsubseteq> 
        snd (THE pair0. pair0 ∈ rclosure nl s2 ∧ fst pair0 = n)"
      from b3 show ?thesis
        apply -
        apply(unfold fclosure_def )
        apply(simp (no_asm) add:Let_def ) thm theI2
        apply(rule_tac ?P="?P1" and a="enttr1"  in theI2)
        apply( blast )
        apply(blast dest:nodeNameisUniInrclosure)
        apply(cut_tac b5)
        apply(rule_tac ?P="?P2" and a="enttr2"  in theI2)
        apply( blast )
        apply(blast dest:nodeNameisUniInrclosure)
        apply(cut_tac b1)
        apply simp
        apply(rule rclosureMono)
        apply assumption+
        apply auto
        
        done
      qed
qed

lemma curFunMono:
  assumes
  a1:"nl:netlists" and
  a2:" s1  \<sqsubseteq>s s2" and
  a3:"( isDefinedIn n nl)"
shows " fclosure nl s1 n \<sqsubseteq>fclosure nl s2 n"
apply -
apply(cut_tac a1 a2 a3)
apply(cut_tac  curFunMono[where nl="nl"])
by(blast )

lemma curFunMonoAny:
assumes a1:"nl:netlists"
shows "ALL s1 s2 n.( ALL n. ( isDefinedIn n nl)--> (s1 n)\<sqsubseteq>  (s2 n)) -->  isDefinedIn n nl --> fclosure nl s1 n \<sqsubseteq>fclosure nl s2 n"
  (is "ALL s1 s2 n. ?P s1 s2 n nl")
proof((rule allI)+,(rule impI)+)
  fix s1 s2 n
  assume b1:"ALL n. isDefinedIn n nl  --> s1 n \<sqsubseteq>  s2 n "  and 
  b2:"isDefinedIn n nl"
  show "fclosure nl s1 n  \<sqsubseteq>  fclosure nl s2 n"
    proof -
      from a1 b1 b2 
      have b20:"( EX! pair. pair:rclosure nl s1 & (fst pair)=n )" (is "EX! enttr. ?P1 enttr ")
        apply -
        apply(drule nodeNameisUniInrclosure )
        apply blast+
        done
      from this obtain enttr1 
        where b3: "?P1  enttr1  &(∀ enttr'. ?P1  enttr'  --> enttr' =enttr1)"
        apply -  apply(erule_tac P="?P1" in  ex1E) apply blast
        done
      from a1 b1 b2
      have b4:"( EX! pair. pair:rclosure nl s2 & (fst pair)=n )" (is "EX! enttr. ?P2 enttr ")
        apply -
        apply(drule nodeNameisUniInrclosure )
        apply blast+
        done
      from this obtain enttr2 
        where b5: "?P2  enttr2  &(∀ enttr'. ?P2  enttr'  --> enttr' =enttr2)"
        apply -  apply(erule_tac P="?P2" in  ex1E) apply blast
        done

      
     (* let et=THE pair0. (pair0):  (rclosure enttrs  s)
                             &(fst pair0) =n in
                            snd  ( et)
      let ?Q=λ  pair0. (pair0):  (rclosure enttrs  s)
                             &(fst pair0) =n in
                            snd  ( et)
      let ?Q="λenttr. (let et = enttr
        in transPart et s1) \<sqsubseteq> 
        (let et =enttr
        in transPart et s2)"*)
      let ?Q="λ enntr. (snd enntr) \<sqsubseteq> 
        snd (THE pair0. pair0 ∈ rclosure nl s2 ∧ fst pair0 = n)"
      from b3 show ?thesis
        apply -
        apply(unfold fclosure_def )
        apply(simp (no_asm) add:Let_def ) thm theI2
        apply(rule_tac ?P="?P1" and a="enttr1"  in theI2)
        apply( blast )
        apply(blast dest:nodeNameisUniInrclosure)
        apply(cut_tac b5)
        apply(rule_tac ?P="?P2" and a="enttr2"  in theI2)
        apply( blast )
        apply(blast dest:nodeNameisUniInrclosure)
        apply(cut_tac b1)
        apply simp
        apply(cut_tac b2,simp)
        apply(rule rclosureMonoAny)
        apply assumption+
        apply auto
        
        done
      qed
qed

lemma curFunMonoAny:
  assumes
  a1:"nl:netlists" and
  a2:" ALL n. ( isDefinedIn n nl)--> (s1 n)\<sqsubseteq>  (s2 n)" and
  a3:"( isDefinedIn n nl)"
shows " fclosure nl s1 n \<sqsubseteq>fclosure nl s2 n"
apply -
apply(cut_tac a1 a2 a3)
apply(cut_tac  curFunMonoAny[where nl="nl"])
by(blast )

lemma isDefinedFanIn:
  assumes a1:"isDefinedIn n nl" and
  a2:"isDelayNames nl n" and
  a3:"isClosed nl" and
  a4:"nl:netlists"
  shows "(fanins ( (lookUp nl n)))~=[]"
proof -
  from a1 a4 have b1: "( 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 -
    apply(rule nodeNameIsUniqueInNet)
    by auto
  then obtain l where b2:"?P l &( ∀ l'. ?P l' --> l'=l)" 
    apply -  apply(erule_tac P="?P " in  ex1E) apply blast
    done
  from a2 have b3:"EX l data.  l=Delay n data & l:nl" (is "EX l data.?Q l data")
    apply -
    apply(unfold isDelayNames_def)
    apply auto
    done
  then obtain l0 data0  where b4:"?Q l0 data0"
    by blast
  then have b5:"?P l0"
    by blast
  with b2  have b6:"l0=l" by blast
  with b5 have b7:"(fanins ( l))=[data0]" (is "?R l")
    apply -
    apply(cut_tac b4)
    apply(auto)
    done
  have "?R (lookUp nl n)"
    apply -
    apply(unfold lookUp_def)
    apply(rule_tac a="l" in  theI2)
    apply(cut_tac b2, blast)
    apply(cut_tac b2,blast)
    apply(subgoal_tac "x=l")
    apply(cut_tac b7,simp)
    apply(cut_tac b2)
    apply blast
    done
  then show ?thesis
    by simp
  qed

lemma hdInSet:"ls ~=[] -->( hd ls ):set ls"
apply(induct_tac ls)
apply auto
done

lemma hdInSet:"ls ~=[] ==> ( hd ls ):set ls"
apply(cut_tac hdInSet[where ls="ls"],blast)
done

lemma latchNameInNet:
  "[|isDelayNames  nlist n; nlist:netlists|]==> isDefinedIn n nlist"
apply(unfold isDelayNames_def isDefinedIn_def)
apply auto
apply(rule_tac x="Delay n data" in bexI)
apply simp
apply simp
done
 
lemma fSeqIsMono:
  assumes
  a1:"nl:netlists" and
  a2:" sq1  \<sqsubseteq>sq  sq2" and
  a3:"isClosed nl" 
shows "ALL n. ( isDefinedIn n nl) --> fSeq nl sq1 t n \<sqsubseteq>  fSeq nl sq2 t n " (is "?P t")
proof(induct_tac t)
  show "?P 0"
    apply -
    apply simp
    apply(cut_tac a1 a2 a3)
    apply(subgoal_tac "(sq1 0)\<sqsubseteq>s(sq2 0)")
    apply(blast intro:curFunMono)
    apply(unfold stateSqLeq_def)
    apply blast
    done
next
  fix t
  assume b1:"?P t"
  show "?P (Suc t)"
    proof(rule allI,rule impI)
      fix n0
      assume b2:"isDefinedIn n0 nl"
      let ?s="(λsq2 x. if isDelayNames nl x
          then let l =  (lookUp nl x); inps = fanins l;
                   v1 = lub (sq2 (Suc t) x) (fSeq nl sq2 t (hd inps))
               in v1
          else sq2 (Suc t) x)"
      show b3:"fSeq nl sq1 (Suc t) n0 \<sqsubseteq>  fSeq nl sq2 (Suc t) n0" (is "?Q")
         proof -
            have "ALL n. ?s sq1 n \<sqsubseteq> ?s sq2 n"
            proof(rule allI)
              fix n1         
              show "?s sq1 n1  \<sqsubseteq> ?s sq2 n1"
                proof(case_tac  "isDelayNames nl n1")
                  assume d1:"isDelayNames nl n1 "
                  from d1 have d5: "isDefinedIn n1 nl"
                    apply -
                    apply(rule  latchNameInNet) 
                    apply assumption+
                    done
                  have d4:"(fanins ( (lookUp nl n1)))~=[]"
                    apply -
                    apply(rule isDefinedFanIn)
                    apply(rule  latchNameInNet)
                    apply assumption+
                    done
                  from a3 this
                  have d3:" isDefinedIn (hd (fanins ( (lookUp nl n1)))) nl"
                    apply -
                    apply(unfold isClosed_def)
                    apply(drule_tac x="n1" in spec)
                     apply(drule_tac x=" hd (fanins ( (lookUp nl n1)))" in spec)
                     apply(cut_tac d5,  simp)
                     apply(subgoal_tac "(hd (fanins ( (lookUp nl n1)))):
                       set (fanins ( (lookUp nl n1)))")
                     apply simp
                     apply(rule hdInSet)
                     by assumption

                  from b1 d3 have d2:"fSeq nl sq1 t (hd (fanins ( (lookUp nl n1)))) \<sqsubseteq> 
                    fSeq nl sq2 t (hd (fanins ( (lookUp nl n1))))"
                    apply -     
                    by blast

                  from d1 d2 show ?thesis
                    apply -                 
                    apply(simp add:Let_def)
                    apply(rule lubMono)
                    apply(cut_tac a2)
                    apply(simp add:stateSqLeq_def stateLeq_def)
                    apply simp
                    done
                  next
                    assume d1:" ¬ isDelayNames nl n1 "
                    from d1 show ?thesis
                       apply -              
                       apply(simp add:Let_def)
                       apply(cut_tac a2)
                       apply(simp add:stateSqLeq_def stateLeq_def)
                       done
                   qed
                 qed
                 then show ?thesis
                   apply -                  
                   apply(simp add:Let_def)
                   apply(rule curFunMono) thm curFunMono 
                   apply(assumption)
                   apply(unfold stateLeq_def)
                   apply blast
                   apply(assumption)
                   done
               qed
             qed
           qed

lemma fSeqIsMono:
  assumes
  a1:"nl:netlists" and
  a2:" sq1  \<sqsubseteq>sq  sq2" and
  a3:"isClosed nl"  and
  a4:" isDefinedIn n nl"
shows  "fSeq nl sq1 t n \<sqsubseteq>  fSeq nl sq2 t n "
  apply(cut_tac fSeqIsMono[where nl="nl" and ?sq1.0="sq1" and ?sq2.0="sq2" and t="t"])
  apply(cut_tac a1 a2 a3 a4, blast)
  apply assumption+
done

lemma fSeqIsMono:
  assumes
  a1:"nl:netlists" and
  a2:" sq1  \<sqsubseteq>sq  sq2" and
  a3:"isClosed nl"  
shows  "fSeq nl sq1 t n \<sqsubseteq>  fSeq nl sq2 t n "
  apply(case_tac "isDefinedIn n nl")
  apply(cut_tac fSeqIsMono[where nl="nl" and ?sq1.0="sq1" and ?sq2.0="sq2" and t="t"])
  
  apply(cut_tac a1 a2 a3 )
  
  apply assumption+
  apply(induct_tac t)
  apply simp
  apply(unfold fclosure_def)
  apply(cut_tac a2,simp)
  apply(unfold stateSqLeq_def stateLeq_def,auto)
  apply(subgoal_tac "¬ isDelayNames nl n")
  apply(simp add:Let_def)
 
  apply(unfold fclosure_def)
  apply auto
  apply(cut_tac a2,unfold stateSqLeq_def stateLeq_def,blast)
  apply(cut_tac a1,blast dest:latchNameInNet)
done

lemma fSeqIsMono1:
  assumes
  a1:"nl:netlists" and
  a2:" sq1  \<sqsubseteq>sq  sq2" and
  a3:"isClosed nl"  
shows  "fSeq nl sq1  \<sqsubseteq>sq  fSeq nl sq2 "
  apply(unfold stateSqLeq_def stateLeq_def)
  apply(rule allI)+
  apply(rule fSeqIsMono)
  apply assumption+
done

lemma fSeqIsMonoAny:
  assumes
  a1:"nl:netlists" and
  a2:" ALL t n . ( isDefinedIn n nl) --> sq1 t n \<sqsubseteq>   sq2 t n" and
  a3:"isClosed nl" 
shows "ALL n. ( isDefinedIn n nl) --> fSeq nl sq1 t n \<sqsubseteq>  fSeq nl sq2 t n " (is "?P t")
proof(induct_tac t)
  show "?P 0"
    apply -
    apply simp
    apply(cut_tac a1 a2 a3)
    apply(subgoal_tac "ALL n.  ( isDefinedIn n nl) -->(sq1 0 n)  \<sqsubseteq> (sq2 0 n)")
    apply(blast intro:curFunMonoAny)
    apply blast
    done
next
  fix t
  assume b1:"?P t"
  show "?P (Suc t)"
    proof(rule allI,rule impI)
      fix n0
      assume b2:"isDefinedIn n0 nl"
      let ?s="(λsq2 x. if isDelayNames nl x
          then let l =  (lookUp nl x); inps = fanins l;
                   v1 = lub (sq2 (Suc t) x) (fSeq nl sq2 t (hd inps))
               in v1
          else sq2 (Suc t) x)"
      show b3:"fSeq nl sq1 (Suc t) n0 \<sqsubseteq>  fSeq nl sq2 (Suc t) n0" (is "?Q")
         proof -
            have "ALL n. isDefinedIn n nl --> ?s sq1 n \<sqsubseteq> ?s sq2 n"
            proof(rule allI,rule impI)
              fix n1        
              assume d6:"isDefinedIn n1 nl"
              show "?s sq1 n1  \<sqsubseteq> ?s sq2 n1"
                proof(case_tac  "isDelayNames nl n1")
                  assume d1:"isDelayNames nl n1 "
                  from d1 have d5: "isDefinedIn n1 nl"
                    apply -
                    apply(rule  latchNameInNet) 
                    apply assumption+
                    done
                  have d4:"(fanins ( (lookUp nl n1)))~=[]"
                    apply -
                    apply(rule isDefinedFanIn)
                    apply(rule  latchNameInNet)
                    apply assumption+
                    done
                  from a3 this
                  have d3:" isDefinedIn (hd (fanins ( (lookUp nl n1)))) nl"
                    apply -
                    apply(unfold isClosed_def)
                    apply(drule_tac x="n1" in spec)
                     apply(drule_tac x=" hd (fanins ( (lookUp nl n1)))" in spec)
                     apply(cut_tac d5,  simp)
                     apply(subgoal_tac "(hd (fanins ( (lookUp nl n1)))):
                       set (fanins ( (lookUp nl n1)))")
                     apply simp
                     apply(rule hdInSet)
                     by assumption

                  from b1 d3 have d2:"fSeq nl sq1 t (hd (fanins ( (lookUp nl n1)))) \<sqsubseteq> 
                    fSeq nl sq2 t (hd (fanins ( (lookUp nl n1))))"
                    apply -     
                    by blast

                  

                  from d1 d2 show ?thesis
                    apply -                 
                    apply(simp add:Let_def)
                    apply(rule lubMono)
                    apply(cut_tac a2) thm latchNameInNet
                    apply(cut_tac a1)
                    apply(blast dest:latchNameInNet)
                    apply simp
                    done
                  next
                    assume d1:" ¬ isDelayNames nl n1 "
                    from d6 d1 show ?thesis
                       apply -              
                       apply(simp add:Let_def)
                       apply(cut_tac a2)
                       apply simp
                       done
                   qed
                 qed
                 then show ?thesis
                   apply -                  
                   apply(simp add:Let_def)
                   apply(rule curFunMonoAny) thm curFunMono 
                   apply(assumption+)
                   
                   done
               qed
             qed
           qed

lemma fSeqIsMonoAny :
  assumes
  a1:"nl:netlists" and
  a2:" ALL t n . ( isDefinedIn n nl) --> sq1 t n \<sqsubseteq>   sq2 t n" and
  a3:"isClosed nl"  and
  a4:" isDefinedIn n nl"
shows  "fSeq nl sq1 t n \<sqsubseteq>  fSeq nl sq2 t n "
  apply(cut_tac fSeqIsMonoAny[where nl="nl" and ?sq1.0="sq1" and ?sq2.0="sq2" and t="t"])
  apply(cut_tac a1 a2 a3 a4, blast)
  apply assumption+
done

lemma fSeqIsMono:
  assumes
  a1:"nl:netlists" and
  a2:" sq1  \<sqsubseteq>sq  sq2" and
  a3:"isClosed nl"  
shows  "fSeq nl sq1 t n \<sqsubseteq>  fSeq nl sq2 t n "
  apply(case_tac "isDefinedIn n nl")
  apply(cut_tac fSeqIsMono[where nl="nl" and ?sq1.0="sq1" and ?sq2.0="sq2" and t="t"])
  
  apply(cut_tac a1 a2 a3 )
  
  apply assumption+
  apply(induct_tac t)
  apply simp
  apply(unfold fclosure_def)
  apply(cut_tac a2,simp)
  apply(unfold stateSqLeq_def stateLeq_def,auto)
  apply(subgoal_tac "¬ isDelayNames nl n")
  apply(simp add:Let_def)
 
  apply(unfold fclosure_def)
  apply auto
  apply(cut_tac a2,unfold stateSqLeq_def stateLeq_def,blast)
  apply(cut_tac a1,blast dest:latchNameInNet)
done

lemma rclosureIsExtensiveAux1:
  assumes
  a1:"nl:netlists" and
  a2:"pair:(rclosure nl s ) "
  shows " ( isDefinedIn n nl)--> 
  (fst pair)=n  --> (s n) \<sqsubseteq> snd pair"
  using a2
  proof(induct,unfold leq_def lub_def ,auto)
  qed
   
lemma fclosureIsExtensive:
  assumes
  a1:"nl:netlists" and
  a2:"isDefinedIn n nl"
  shows 
  "(s n) \<sqsubseteq> fclosure nl s n "
  proof(unfold fclosure_def, cut_tac a1, simp add:Let_def,rule conjI,rule impI)
    show "s n \<sqsubseteq>  snd (THE pair0. pair0 ∈ rclosure nl s ∧ fst pair0 = n)"
      (is "?Q (THE pair0. pair0 ∈ rclosure nl s ∧ fst pair0 = n)")
      proof -
        let ?P="λ pair0.  pair0 ∈ rclosure nl s ∧ fst pair0 = n"
        have c1:"EX! pair0. ?P pair0" thm pairIsUniInrclosure
         apply (cut_tac a1 ,rule nodeNameisUniInrclosure)
         by assumption+
       then obtain pair1 where c2:"?P pair1 &(ALL pair.?P pair --> pair=pair1)"
         by blast
       show ?thesis
         apply(rule_tac a="pair1" in theI2)
         apply(cut_tac c2,simp)
         apply(cut_tac c2)
         apply blast
         apply(cut_tac a1 a2)
         apply(blast dest:rclosureIsExtensiveAux1)
         done
       qed
     next
       from a1 show " ¬ isDefinedIn n nl --> s n \<sqsubseteq>  s n"
         by(cut_tac a2,simp)
     qed


lemma fclosureIsExtensive1:
  assumes
  a1:"nl:netlists"
  shows 
  "(s n) \<sqsubseteq> fclosure nl s n"
  apply(case_tac "isDefinedIn n nl")
  apply(cut_tac a1,blast dest:fclosureIsExtensive)
  apply(simp add: fclosure_def)
  apply(rule leqReflexive)
  done
lemma lubSelf: " a \<sqsubseteq> (lub a  b )"
  apply(unfold leq_def lub_def)
  apply auto
done

lemma lubComm:" lub a b= lub b a"
  apply(unfold leq_def lub_def)
  apply auto
done

lemma leqIsTrans:
  "[|a \<sqsubseteq> b; b \<sqsubseteq> c|]==> a\<sqsubseteq> c" 
  apply(unfold leq_def lub_def) 
  apply(simp add:Pair_fst_snd_eq)
  apply auto
  done

lemma sqLeqIsTrans:
  "[|a \<sqsubseteq>sq  b; b \<sqsubseteq>sq  c|]==> a\<sqsubseteq>sq  c" 
  apply(unfold stateSqLeq_def stateLeq_def ) 
  apply(auto dest:leqIsTrans)
  apply(blast dest:leqIsTrans)
  done

lemma fSeqIsExtensive1:
  assumes
  a1:"nl:netlists" and a2:"isClosed nl"
  shows 
  "ALL n. sq t n \<sqsubseteq> fSeq nl sq t n"
  proof(induct_tac t)
    
    show "ALL n. sq 0 n \<sqsubseteq>  fSeq nl sq 0 n"
      apply(rule allI)
      apply(cut_tac a1)
      apply simp
      apply(blast dest:fclosureIsExtensive1)+
      done
  next
    fix na
    assume b1:"ALL n. sq na n \<sqsubseteq>  fSeq nl sq na n"
    show "ALL n. sq (Suc na) n \<sqsubseteq>  fSeq nl sq (Suc na) n"
    proof(rule allI)
      fix n
      have "isDelayNames nl n | ~ isDelayNames nl n "
        by simp
      moreover
      {
        assume d1:"isDelayNames nl n "
        from d1 have d5: "isDefinedIn n nl"
          apply -
          apply(rule  latchNameInNet) 
          apply assumption+
          done
        have d4:"(fanins ( (lookUp nl n)))~=[]"
          apply -
          apply(rule isDefinedFanIn)
          apply(rule  latchNameInNet)
          apply assumption+
          done
        from a2 this
        have d3:" isDefinedIn (hd (fanins ( (lookUp nl n)))) nl"
          apply -
          apply(unfold isClosed_def)
          apply(drule_tac x="n" in spec)
          apply(drule_tac x=" hd (fanins ( (lookUp nl n)))" in spec)
          apply(cut_tac d5,  simp)
          apply(subgoal_tac "(hd (fanins ( (lookUp nl n)))):
            set (fanins ( (lookUp nl n)))")
          apply simp
          apply(rule hdInSet)
          by assumption
        from b1 d3 have d2:" sq na (hd (fanins ( (lookUp nl n)))) \<sqsubseteq> 
          fSeq nl sq na (hd (fanins ( (lookUp nl n))))"
          apply -       
          by blast

        let ?s="(λx. if isDelayNames nl x
          then let l = lookUp nl x; inps = fanins l;
          v1 = lub (sq (Suc na) x) (fSeq nl sq na (hd inps))
          in v1
          else sq (Suc na) x)"
        have "(sq (Suc na) n ) \<sqsubseteq> ?s n"
          apply(cut_tac d1,simp add:Let_def)
          apply(rule lubSelf)
          done
        also have d1:"… \<sqsubseteq>  fclosure nl ?s n" 
          apply(rule  fclosureIsExtensive1)
          by assumption
        ultimately have d0:"(sq (Suc na) n ) \<sqsubseteq>  fclosure nl ?s n" 
          by (auto dest:leqIsTrans)
        have " sq (Suc na) n \<sqsubseteq>  fSeq nl sq (Suc na) n"        
          apply(simp add:Let_def)+
          apply(cut_tac d0)
          apply simp
          done
      }
      moreover
      {
        assume d1:" ¬ isDelayNames nl n "
        let ?s="(λx. if isDelayNames nl x
          then let l = lookUp nl x; inps = fanins l;
          v1 = lub (sq (Suc na) x) (fSeq nl sq na (hd inps))
          in v1
          else sq (Suc na) x)"
        have "(sq (Suc na) n ) \<sqsubseteq> ?s n"
          apply(cut_tac d1,simp add:Let_def)
          apply(rule leqReflexive)
          done
        also have d1:"… \<sqsubseteq>  fclosure nl ?s n" 
          apply(rule  fclosureIsExtensive1)
          by assumption
        ultimately have d0:"(sq (Suc na) n ) \<sqsubseteq>  fclosure nl ?s n" 
          by (auto dest:leqIsTrans)
        from d1 have  " sq (Suc na) n \<sqsubseteq>  fSeq nl sq (Suc na) n"
          apply -                   
          apply(simp add:Let_def)
          apply(cut_tac d0)
          apply simp
          done
        }
        ultimately show " sq (Suc na) n \<sqsubseteq>  fSeq nl sq (Suc na) n"
          by blast
      qed
    qed
          
lemma fSeqIsExtensive:
  assumes
  a1:"nl:netlists" and a2:"isClosed nl"
  shows 
  "sq   \<sqsubseteq>sq  fSeq nl sq "  
  apply(cut_tac a1 a2)
  
  apply(unfold stateSqLeq_def stateLeq_def) 
  apply(rule allI)
  apply(rule fSeqIsExtensive1)
  by assumption+


lemma lookUpIsTheDev:
  assumes
  a1:"nl:netlists" and
  a2:"
  enttr ∈ nl ∧ (enttr = Input n ∨ (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab))"
  shows "lookUp nl n=enttr "
proof -
  have b1:"EX! enttr. 
  enttr ∈ nl ∧ (enttr = Input n ∨ (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab))"
    apply(rule nodeNameIsUniqueInNet)
    apply assumption+
    apply(unfold isDefinedIn_def)
    apply(cut_tac a2,auto)
    apply(rule_tac x="Input n" in bexI)
    apply simp+
    apply(rule_tac x=" Delay n data " in bexI)
    apply simp+
    apply(rule_tac x=" Gate n inps tab " in bexI)
    apply simp+
    done
  let ?R="λ  enttr.
                 enttr ∈ nl ∧
                 (enttr = Input n ∨ (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab))"
  from b1 obtain l where b2:"?R  l  &(∀ l0. ?R  l0 --> l0=l)" by blast
  from a2 this  have "l=enttr" by auto
  with b2  have "?R  enttr  &(∀ l0. ?R  l0 --> l0=enttr)" by auto
  then show ?thesis
    apply -
    apply(unfold lookUp_def)
    apply(simp add:Let_def)
    apply(rule_tac a="enttr" in theI2) thm theI
    apply blast
    apply blast
    apply blast
    done
qed
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

end

lemma nodeNameIsUniqueInNet:

  nl ∈ netlists ==>
  ∀n. isDefinedIn n nl -->
      (∃!enttr.
          enttrnl ∧
          (enttr = Input n ∨
           (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab)))

lemma nodeNameIsUniqueInNet:

  [|nl ∈ netlists; isDefinedIn n nl|]
  ==> ∃!enttr.
         enttrnl ∧
         (enttr = Input n ∨
          (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab))

lemma nodeNameIsUniqueInNet2:

  [|nl ∈ netlists; isDefinedIn n nl|] ==> ∃!enttr. enttrnl ∧ fanout enttr = n

lemma InputNameDelayNameIsDiff:

  [|nl ∈ netlists; Input nnl; Delay n datanl|] ==> False

lemma InputNameGateNameIsDiff:

  [|nl ∈ netlists; Input nnl; Gate n inps tabnl|] ==> False

lemma DelayNameGateNameIsDiff:

  [|nl ∈ netlists; Delay n datanl; Gate n inps tabnl|] ==> False

lemma DelayNameIsSame:

  [|nl ∈ netlists; Delay n datanl; Delay n data0.0nl|] ==> data = data0.0

lemma GateNameIsSame:

  [|nl ∈ netlists; Gate n inps tabnl; Gate n inps1.0 tab1.0nl|]
  ==> inps = inps1.0tab = tab1.0

lemma stTransOfNil:

  rclosure {} s = {}

lemma stateListLeq:

stateLs1 inps.
     length stateLs1 = length inps -->
     length stateLs = length inps -->
     (∀pair0.
         pair0 mem zip inps stateLs -->
         pair0 ∈ rclosure nl s1.0 ∧
         (∀pair2.
             pair2 ∈ rclosure nl s2.0 ∧ fst pair0 = fst pair2 -->
             snd pair0 \<sqsubseteq>  snd pair2)) -->
     (∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s2.0) -->
     stateSqLeq1 stateLs stateLs1

lemma stateListLeq:

  [|length stateLs1.0 = length inps; length stateLs = length inps;
   ∀pair0.
      pair0 mem zip inps stateLs -->
      pair0 ∈ rclosure nl s1.0 ∧
      (∀pair2.
          pair2 ∈ rclosure nl s2.0 ∧ fst pair0 = fst pair2 -->
          snd pair0 \<sqsubseteq>  snd pair2);
   ∀pair0. pair0 mem zip inps stateLs1.0 --> pair0 ∈ rclosure nl s2.0|]
  ==> stateSqLeq1 stateLs stateLs1.0

lemma ftabIsMono:

  [|∀pair0.
       pair0 mem zip inps stateLs -->
       pair0 ∈ rclosure nl s1.0 ∧
       (∀pair2.
           pair2 ∈ rclosure nl s2.0 -->
           fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2);
   length stateLs = length inps;
   ∀pair0. pair0 mem zip inps stateLsa --> pair0 ∈ rclosure nl s2.0;
   length stateLsa = length inps|]
  ==> (∀l. l ∈ set tab --> length l = length inps) -->
      funOfTab tab stateLs \<sqsubseteq>  funOfTab tab stateLsa

lemma ftabIsMono:

  [|∀pair0.
       pair0 mem zip inps stateLs -->
       pair0 ∈ rclosure nl s1.0 ∧
       (∀pair2.
           pair2 ∈ rclosure nl s2.0 -->
           fst pair0 = fst pair2 --> snd pair0 \<sqsubseteq>  snd pair2);
   length stateLs = length inps;
   ∀pair0. pair0 mem zip inps stateLsa --> pair0 ∈ rclosure nl s2.0;
   length stateLsa = length inps; ∀l. l ∈ set tab --> length l = length inps|]
  ==> funOfTab tab stateLs \<sqsubseteq>  funOfTab tab stateLsa

lemma rclosureMono:

  nl ∈ netlists ==>
  ∀s1 s2.
     s1 \<sqsubseteq>s s2 -->
     pair1.0 ∈ rclosure nl s1 -->
     (∀pair2.
         pair2 ∈ rclosure nl s2 -->
         fst pair1.0 = fst pair2 --> snd pair1.0 \<sqsubseteq>  snd pair2)

lemma rclosureMonoAny:

  nl ∈ netlists ==>
  ∀s1 s2.
     (∀n. isDefinedIn n nl --> s1 n \<sqsubseteq>  s2 n) -->
     pair1.0 ∈ rclosure nl s1 -->
     (∀pair2.
         pair2 ∈ rclosure nl s2 -->
         fst pair1.0 = fst pair2 --> snd pair1.0 \<sqsubseteq>  snd pair2)

lemma rclosureMono:

  [|nl ∈ netlists; s1.0 \<sqsubseteq>s s2.0; pair1.0 ∈ rclosure nl s1.0;
   pair2.0 ∈ rclosure nl s2.0; fst pair1.0 = fst pair2.0|]
  ==> snd pair1.0 \<sqsubseteq>  snd pair2.0

lemma rclosureMonoAny:

  [|nl ∈ netlists; ∀n. isDefinedIn n nl --> s1.0 n \<sqsubseteq>  s2.0 n;
   pair1.0 ∈ rclosure nl s1.0; pair2.0 ∈ rclosure nl s2.0;
   fst pair1.0 = fst pair2.0|]
  ==> snd pair1.0 \<sqsubseteq>  snd pair2.0

lemma inputNodeIsUniInFclosure:

  [|pair0.0 ∈ rclosure nl s; Input nnl; nl ∈ netlists|]
  ==> fst pair0.0 = n --> pair0.0 = (n, s n)

lemma latchNodeIsUniInFclosure:

  [|pair0.0 ∈ rclosure nl s; Delay n datanl; nl ∈ netlists|]
  ==> fst pair0.0 = n --> pair0.0 = (n, s n)

lemma stateListEq:

stateLs1 inps.
     length stateLs1 = length inps -->
     length stateLs = length inps -->
     (∀pair0.
         pair0 mem zip inps stateLs -->
         pair0 ∈ rclosure nl s ∧
         (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0)) -->
     (∀pair0. pair0 mem zip inps stateLs1 --> pair0 ∈ rclosure nl s) -->
     stateLs = stateLs1

lemma stateListEq:

  [|length stateLs1.0 = length inps; length stateLs = length inps;
   ∀pair0.
      pair0 mem zip inps stateLs -->
      pair0 ∈ rclosure nl s ∧
      (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0);
   ∀pair0. pair0 mem zip inps stateLs1.0 --> pair0 ∈ rclosure nl s|]
  ==> stateLs = stateLs1.0

lemma GateNodeIsUniInFclosure:

  [|pair0.0 ∈ rclosure nl s; length stateLs = length inps; Gate n inps tabnl;
   nl ∈ netlists;
   ∀pair0.
      pair0 mem zip inps stateLs -->
      pair0 ∈ rclosure nl s ∧
      (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0)|]
  ==> fst pair0.0 = n --> pair0.0 = (n, lub (funOfTab tab stateLs) (s n))

lemma GateNodeIsUniInFclosure:

  [|pair0.0 ∈ rclosure nl s; length stateLs = length inps; Gate n inps tabnl;
   nl ∈ netlists;
   ∀pair0.
      pair0 mem zip inps stateLs -->
      pair0 ∈ rclosure nl s ∧
      (∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0);
   fst pair0.0 = n|]
  ==> pair0.0 = (n, lub (funOfTab tab stateLs) (s n))

lemma pairIsUniInrclosure:

  [|nl ∈ netlists; pair0.0 ∈ rclosure nl s|]
  ==> ∃!pair. pair ∈ rclosure nl s ∧ fst pair = fst pair0.0

lemma stClosureMono:

  [|x ∈ rclosure nl s; nlnl0.0|] ==> x ∈ rclosure nl0.0 s

lemma stClosureMono1:

  nlnl0.0 ==> rclosure nl s ⊆ rclosure nl0.0 s

lemma aux1:

  (∀n. n mem inps --> (∃pair. pair ∈ rclosure nl s ∧ fst pair = n)) -->
  pair0.0 mem
  zip inps
   (map (λn. snd (SOME pair. pair ∈ rclosure nl s ∧ fst pair = n)) inps) -->
  pair0.0 ∈ rclosure nl s

lemma aux1:

  [|∀n. n mem inps --> (∃pair. pair ∈ rclosure nl s ∧ fst pair = n);
   pair0.0 mem
   zip inps
    (map (λn. snd (SOME pair. pair ∈ rclosure nl s ∧ fst pair = n)) inps)|]
  ==> pair0.0 ∈ rclosure nl s

lemma nodeNameisUniInrclosure1:

  nl ∈ netlists ==>
  ∀n. isDefinedIn n nl --> (∃pair. pair ∈ rclosure nl s ∧ fst pair = n)

lemma nodeNameisUniInrclosure:

  [|nl ∈ netlists; isDefinedIn n nl|]
  ==> ∃!pair. pair ∈ rclosure nl s ∧ fst pair = n

lemma curFunMono:

  nl ∈ netlists ==>
  ∀s1 s2 n.
     s1 \<sqsubseteq>s s2 -->
     isDefinedIn n nl --> fclosure nl s1 n \<sqsubseteq>  fclosure nl s2 n

lemma curFunMono:

  [|nl ∈ netlists; s1.0 \<sqsubseteq>s s2.0; isDefinedIn n nl|]
  ==> fclosure nl s1.0 n \<sqsubseteq>  fclosure nl s2.0 n

lemma curFunMonoAny:

  nl ∈ netlists ==>
  ∀s1 s2 n.
     (∀n. isDefinedIn n nl --> s1 n \<sqsubseteq>  s2 n) -->
     isDefinedIn n nl --> fclosure nl s1 n \<sqsubseteq>  fclosure nl s2 n

lemma curFunMonoAny:

  [|nl ∈ netlists; ∀n. isDefinedIn n nl --> s1.0 n \<sqsubseteq>  s2.0 n;
   isDefinedIn n nl|]
  ==> fclosure nl s1.0 n \<sqsubseteq>  fclosure nl s2.0 n

lemma isDefinedFanIn:

  [|isDefinedIn n nl; isDelayNames nl n; isClosed nl; nl ∈ netlists|]
  ==> fanins (lookUp nl n) ≠ []

lemma hdInSet:

  ls ≠ [] --> hd ls ∈ set ls

lemma hdInSet:

  ls ≠ [] ==> hd ls ∈ set ls

lemma latchNameInNet:

  [|isDelayNames nlist n; nlist ∈ netlists|] ==> isDefinedIn n nlist

lemma fSeqIsMono:

  [|nl ∈ netlists; sq1.0 \<sqsubseteq>sq  sq2.0; isClosed nl|]
  ==> ∀n. isDefinedIn n nl --> fSeq nl sq1.0 t n \<sqsubseteq>  fSeq nl sq2.0 t n

lemma fSeqIsMono:

  [|nl ∈ netlists; sq1.0 \<sqsubseteq>sq  sq2.0; isClosed nl; isDefinedIn n nl|]
  ==> fSeq nl sq1.0 t n \<sqsubseteq>  fSeq nl sq2.0 t n

lemma fSeqIsMono:

  [|nl ∈ netlists; sq1.0 \<sqsubseteq>sq  sq2.0; isClosed nl|]
  ==> fSeq nl sq1.0 t n \<sqsubseteq>  fSeq nl sq2.0 t n

lemma fSeqIsMono1:

  [|nl ∈ netlists; sq1.0 \<sqsubseteq>sq  sq2.0; isClosed nl|]
  ==> fSeq nl sq1.0 \<sqsubseteq>sq  fSeq nl sq2.0

lemma fSeqIsMonoAny:

  [|nl ∈ netlists; ∀t n. isDefinedIn n nl --> sq1.0 t n \<sqsubseteq>  sq2.0 t n;
   isClosed nl|]
  ==> ∀n. isDefinedIn n nl --> fSeq nl sq1.0 t n \<sqsubseteq>  fSeq nl sq2.0 t n

lemma fSeqIsMonoAny:

  [|nl ∈ netlists; ∀t n. isDefinedIn n nl --> sq1.0 t n \<sqsubseteq>  sq2.0 t n;
   isClosed nl; isDefinedIn n nl|]
  ==> fSeq nl sq1.0 t n \<sqsubseteq>  fSeq nl sq2.0 t n

lemma fSeqIsMono:

  [|nl ∈ netlists; sq1.0 \<sqsubseteq>sq  sq2.0; isClosed nl|]
  ==> fSeq nl sq1.0 t n \<sqsubseteq>  fSeq nl sq2.0 t n

lemma rclosureIsExtensiveAux1:

  [|nl ∈ netlists; pair ∈ rclosure nl s|]
  ==> isDefinedIn n nl --> fst pair = n --> s n \<sqsubseteq>  snd pair

lemma fclosureIsExtensive:

  [|nl ∈ netlists; isDefinedIn n nl|] ==> s n \<sqsubseteq>  fclosure nl s n

lemma fclosureIsExtensive1:

  nl ∈ netlists ==> s n \<sqsubseteq>  fclosure nl s n

lemma lubSelf:

  a \<sqsubseteq>  lub a b

lemma lubComm:

  lub a b = lub b a

lemma leqIsTrans:

  [|a \<sqsubseteq>  b; b \<sqsubseteq>  c|] ==> a \<sqsubseteq>  c

lemma sqLeqIsTrans:

  [|a \<sqsubseteq>sq  b; b \<sqsubseteq>sq  c|] ==> a \<sqsubseteq>sq  c

lemma fSeqIsExtensive1:

  [|nl ∈ netlists; isClosed nl|] ==> ∀n. sq t n \<sqsubseteq>  fSeq nl sq t n

lemma fSeqIsExtensive:

  [|nl ∈ netlists; isClosed nl|] ==> sq \<sqsubseteq>sq  fSeq nl sq

lemma lookUpIsTheDev:

  [|nl ∈ netlists;
   enttrnl ∧
   (enttr = Input n ∨
    (∃data. enttr = Delay n data) ∨ (∃inps tab. enttr = Gate n inps tab))|]
  ==> lookUp nl n = enttr

lemma pairInzipList:

stateLs. pair mem zip inps stateLs --> fst pair mem inps

lemma pairInzipList:

  pair mem zip inps stateLs ==> fst pair mem inps