Theory subsetNet

Up to index of Isabelle/HOL/net

theory subsetNet
imports net
begin

theory subsetNet=net:

consts subNetList::" entity set =>  entity set => entity set"
inductive "subNetList nl snl"
intros
  subAddself:
  "[|enttr: snl;enttr: nl  |]==> enttr:subNetList nl snl"
  subAddLink:
  "[|enttr0: subNetList nl snl;  enttr1:nl;
  (fanout ( enttr1)):(set (fanins ( enttr0))) |]==> enttr1:subNetList nl snl"

lemma subsetNetRef:
  "x:subNetList nl nl'==> x: nl"
apply(erule subNetList.induct)
apply auto
done

lemma subsetNetlistNil:"x:subNetList {} snl ==> False"
apply(erule subNetList.induct)
apply auto
done

lemma subsetNetUn1:
assumes a1:"x: subNetList ( nl) (snl1 Un snl2)"
shows "x:subNetList ( nl) (snl1) | x:subNetList ( nl) (snl2)"
apply(cut_tac a1)
apply(erule subNetList.induct)
apply(blast dest:subAddself)
apply(blast dest:subAddLink)
done

lemma subsetUnRev1:
assumes 
a1:"x:subNetList ( nl) (snl1) | x:subNetList ( nl) (snl2)"
shows "x: subNetList ( nl) (snl1 Un snl2)"

apply(cut_tac a1)
apply(erule disjE)
apply(erule subNetList.induct)
apply(blast intro:subAddself)
apply(blast intro:subAddLink)
apply(erule subNetList.induct)
apply(blast intro:subAddself)
apply(blast intro:subAddLink)
done
lemma subsetNetUn:
" subNetList ( nl) (snl1 Un snl2) =  subNetList ( nl) (snl1) Un  subNetList ( nl) (snl2)"
apply(rule equalityI)
apply(rule subsetI)
apply(drule subsetNetUn1)
apply blast
apply(rule subsetI)
apply(erule UnE)
apply( blast intro:subsetUnRev1)+
done

lemma subsetNetAdd:
assumes a1:"x: subNetList ({l}Un nl) (snl)" 
  
  
shows "x:subNetList ( nl) (snl) | x=l | x:subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( l)))} "
apply(cut_tac a1)
apply(erule subNetList.induct)
apply(blast dest:subAddself)
apply(erule UnE)
apply simp
apply(erule disjE)
apply(blast dest:subAddLink)
apply(erule disjE)
apply(rule disjI2)+
apply(blast intro:subAddself)
apply(blast dest:subAddLink)
done

lemma subsetNetAux:"[| x:subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( l)))};
  l: subNetList ({l} ∪ nl) snl|] 
  ==> x: subNetList ({l}Un nl) (snl)"

apply(erule subNetList.induct)
apply(blast dest:subAddLink)
apply(blast dest:subAddLink)
done

lemma subsetNetMono:"[|x:subNetList M nl ;M ⊆ N|]==> x:subNetList N nl"

apply(erule subNetList.induct)
apply(blast intro:subAddself)
apply(blast intro:subAddLink)
done

lemma subsetNetMono2:"[|x:subNetList M nl ;nl ⊆ nl'|]==> x:subNetList M nl'"

apply(erule subNetList.induct)
apply(blast intro:subAddself)
apply(blast intro:subAddLink)
done


lemma subsetNet1:"[| l: subNetList ({l} ∪ nl) snl|]==> 
subNetList ({l}Un nl) (snl) =
(subNetList ( nl) (snl)) Un {l} Un subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( l)))}"
apply(rule equalityI)
apply(rule subsetI)
apply(rotate_tac 1,drule subsetNetAdd)
apply blast
apply(rule subsetI)
apply(erule UnE)+
apply(rule subsetNetMono,blast)
apply blast
apply simp
apply(blast intro:subsetNetAux)
done

lemma subsetNetAux1: "[|~ l: subNetList ({l} ∪ nl) snl;x: subNetList ({l}Un nl) (snl)|] ==>x: subNetList ( nl) (snl)"
  apply(erule subNetList.induct)
  apply(erule UnE)
  apply simp thm subAddself
  apply(drule_tac enttr="l" and nl="{l}Un nl" in subAddself)
  apply simp
  apply simp
  apply(drule  subAddself)
  apply simp apply simp
  apply(blast dest:  subAddLink)
done

lemma subsetNet2: "[|~ l: subNetList ({l} ∪ nl) snl|] ==>subNetList ({l}Un nl) (snl)=subNetList ( nl) (snl)"
apply(rule equalityI)
apply(rule subsetI)
apply(drule subsetNetAux1)
apply blast+
apply(rule subsetI)
apply(rule subsetNetMono,blast)
apply blast
done

lemma subsetNetlistNil:"subNetList {} snl={}"
apply(blast dest:subsetNetlistNil)
done

lemma subsetNetJoint:"x:subNetList nl nl' ==> x:subNetList nl (nl' Int nl)"
apply(erule subNetList.induct)
apply(blast intro:subAddself)
apply(blast intro:subAddLink)
done

lemma subsetNetJoint1:"subNetList nl nl'=subNetList nl (nl' Int nl)"
apply(rule equalityI)
apply(rule subsetI)
apply(drule subsetNetJoint)
apply blast
apply(rule subsetI)
apply(rule subsetNetMono2,blast)
apply blast
done



lemma subNetIsNetlist:
assumes
  a1:"nl:netlists"
  shows  
  " (ALL snl. snl ⊆ nl --> (subNetList nl snl):netlists)"
  (is  "?P nl")
using a1
proof(induct)
show "?P {}"   
  apply(rule allI,rule impI)
  apply(simp add:subsetNetlistNil)
  apply( auto intro:nilNetList)
  done

next
  fix n nl
  assume b1:"¬ isDefinedIn n nl" and
    b2:"nl ∈ netlists" and
    b3:"?P nl"
  let ?l="(Input n)"
  
  let ?nl="{?l} ∪ nl "
  
  show "?P ?nl" (is " (ALL snl. ?P2 snl --> ?P3 snl)") 
  proof(rule allI,rule impI)
    fix snl
    assume 
    c2:"snl ⊆ ?nl" 
    show " ?P3 snl"
    proof -
      have "?l: subNetList ({?l} ∪ nl) snl | ?l ~: subNetList ({?l} ∪ nl) snl"
        by blast
      moreover
      {assume 
        d1:"?l: subNetList ({?l} ∪ nl) snl"
        from d1 have d2:" subNetList ({?l} ∪ nl) snl=
          (subNetList ( nl) (snl)) Un {?l} Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}"
          by(rule  subsetNet1)
        have "(subNetList ( nl) (snl))=(subNetList ( nl) (snl Int nl))"
          by (rule  subsetNetJoint1)
        with d2 have d3:"subNetList ({?l} ∪ nl) snl=
          (subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))} Un {?l}"
          by simp

        let ?S0="((snl Int nl) Un {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))})"
        let ?S1="subNetList ( nl) ?S0"

        have d4:"(subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}=?S1 "
          thm subsetNetUn
          apply(simp add: subsetNetUn[THEN sym])+
          thm subsetNetUn[THEN sym]
          
          done

        have d5:"?S1:netlists"
          apply(cut_tac b3)
          apply(drule_tac x="?S0" in spec)
          apply blast
          done

        let ?S2="(subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}"
        from d5 d4  have "?S2 :netlists"
          by simp

        
        have " {?l} Un ?S1 :netlists"
          apply(rule addInput)
          apply(rule ccontr)
          apply(simp)
          apply(subgoal_tac "?S1 ⊆  nl")
          apply(subgoal_tac "isDefinedIn n nl")
          apply(cut_tac b1, simp)
          prefer 2
          apply(blast intro: subsetNetRef)
          prefer 2
          apply(cut_tac d5,assumption)
          apply(subgoal_tac "(subNetList nl (snl ∩ nl)) ⊆ nl")
          apply(unfold isDefinedIn_def,blast)
          apply(blast dest: subsetNetRef)
          done

 
        with d4  have "{?l} Un ?S2 :netlists" by auto
        with d2 d3  have ?thesis by auto
      }
      moreover 
      {assume 
        d1:"?l~: subNetList ({?l} ∪ nl) snl"
        have "subNetList ({?l}Un nl) (snl)=subNetList ( nl) (snl)"
          apply(rule subsetNet2)
          apply(cut_tac d1)
          by assumption
        also have "…=subNetList ( nl) (snl Int nl)"
          by (rule  subsetNetJoint1)
        also have "subNetList ( nl) (snl Int nl):netlists"
          apply(cut_tac b3)
          apply(drule_tac x="?S0" in spec)
          apply blast
          done
        ultimately have ?thesis by simp
      }
      ultimately show ?thesis by blast
    qed
  qed
next
  fix data n nl
  assume b1:"¬ isDefinedIn n nl" and
    b2:"nl ∈ netlists" and
    b3:"?P nl"
  let ?l="(Delay n data)"  
  let ?nl="({?l} ∪ nl) "
  show "?P ?nl" (is " (ALL snl. ?P2 snl --> ?P3 snl)") 
  proof(rule allI,rule impI)
    fix snl
    assume 
    c2:"snl ⊆ ?nl" 
    show " ?P3 snl"
    proof -
      have "?l: subNetList ({?l} ∪ nl) snl | ?l ~: subNetList ({?l} ∪ nl) snl"
        by blast
      moreover
      {assume 
        d1:"?l: subNetList ({?l} ∪ nl) snl"
        from d1 have d2:" subNetList ({?l} ∪ nl) snl=
          (subNetList ( nl) (snl)) Un {?l} Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}"
          by(rule  subsetNet1)
        have "(subNetList ( nl) (snl))=(subNetList ( nl) (snl Int nl))"
          by (rule  subsetNetJoint1)
        with d2 have d3:"subNetList ({?l} ∪ nl) snl=
          (subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))} Un {?l}"
          by simp

        let ?S0="((snl Int nl) Un {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))})"
        let ?S1="subNetList ( nl) ?S0"

        have d4:"(subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}=?S1 "
          apply(simp add: subsetNetUn)
          done

        have d5:"?S1:netlists"
          apply(cut_tac b3)
          apply(drule_tac x="?S0" in spec)
          apply blast
          done

        let ?S2="(subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}"
        from d5 d4  have "?S2 :netlists"
          by simp

        
        have " {?l} Un ?S1 :netlists"
          apply(rule addDelay)
          apply(rule ccontr)
          apply(simp)
          apply(subgoal_tac "?S1 ⊆  nl")
          apply(subgoal_tac "isDefinedIn n nl")
          apply(cut_tac b1, simp)
          apply(unfold isDefinedIn_def, simp add:subsetNetUn[THEN sym])
          thm subsetNetRef
          apply(blast intro: subsetNetRef)+
          apply(cut_tac d5,assumption)
          done

 
        with d4  have "{?l} Un ?S2 :netlists" by auto
        with d2 d3  have ?thesis by auto
      }
      moreover 
      {assume 
        d1:"?l~: subNetList ({?l} ∪ nl) snl"
        have "subNetList ({?l}Un nl) (snl)=subNetList ( nl) (snl)"
          apply(rule subsetNet2)
          apply(cut_tac d1)
          by assumption
        also have "…=subNetList ( nl) (snl Int nl)"
          by (rule  subsetNetJoint1)
        also have "subNetList ( nl) (snl Int nl):netlists"
          apply(cut_tac b3)
          apply(drule_tac x="?S0" in spec)
          apply blast
          done
        ultimately have ?thesis by simp
      }
      ultimately show ?thesis by blast
    qed
  qed
next
  fix inps n nl tab 
  assume b1:"¬ isDefinedIn n nl" and
    b2:"nl ∈ netlists" and
    b3:"?P nl" 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 n inps (tab))"  
  let ?nl="({?l} ∪ nl) "
  show "?P ?nl" (is " (ALL snl. ?P2 snl --> ?P3 snl)") 
  proof(rule allI,rule impI)
    fix snl
    assume 
    c2:"snl ⊆ ?nl" 
    show " ?P3 snl"
    proof -
      have "?l: subNetList ({?l} ∪ nl) snl | ?l ~: subNetList ({?l} ∪ nl) snl"
        by blast
      moreover
      {assume 
        d1:"?l: subNetList ({?l} ∪ nl) snl"
        from d1 have d2:" subNetList ({?l} ∪ nl) snl=
          (subNetList ( nl) (snl)) Un {?l} Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}"
          by(rule  subsetNet1)
        have "(subNetList ( nl) (snl))=(subNetList ( nl) (snl Int nl))"
          by (rule  subsetNetJoint1)
        with d2 have d3:"subNetList ({?l} ∪ nl) snl=
          (subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))} Un {?l}"
          by simp

        let ?S0="((snl Int nl) Un {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))})"
        let ?S1="subNetList ( nl) ?S0"

        have d4:"(subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}=?S1 "
          apply(simp add: subsetNetUn)
          done

        have d5:"?S1:netlists"
          apply(cut_tac b3)
          apply(drule_tac x="?S0" in spec)
          apply blast
          done

        let ?S2="(subNetList ( nl) (snl Int nl)) Un 
          subNetList ( nl) {l0. l0:nl&(fanout ( l0)):(set (fanins ( ?l)))}"
        from d5 d4  have "?S2 :netlists"
          by simp

        let ?R="( λ n  enttr.(enttr):nl 
                             &(( enttr)= Input n |
                               ( EX data .( enttr)=Delay n data )|
                               (EX inps tab.( enttr)= Gate n inps tab)))"
        
        have d6:"∀n. n mem inps --> isDefinedIn n ?S1"
          proof(rule allI,rule impI)
            fix na
            assume e1:"na mem inps"
            show "isDefinedIn na ?S1"
            proof -
              have "EX! x. ?R na x"
                apply(rule nodeNameIsUniqueInNet)
                apply assumption+
                apply(cut_tac b5 e1, blast)
                done
              then obtain l where e2:"?R na l & (ALL l0. ?R na l0 --> l0=l)"
                by blast
              from e1 have e3:"na : set inps"
                by (simp add:mem_iff)
              from e2 e3 have "  ∃l∈subNetList nl {l0 ∈ nl. fanout l0 ∈ set inps}.
              fanout l = na"
                apply -
                apply(rule_tac x="THE x.?R na x" in bexI)
                apply(rule_tac a="l" in theI2)
                apply blast
                apply blast
                prefer 2
                apply(rule_tac a="l" in theI2)
                apply blast
                apply blast
                apply(rule  subAddself)
                apply( auto)
                done
              then show ?thesis   
                apply(unfold isDefinedIn_def)
                apply(cut_tac d4[THEN sym])
                apply simp
                apply blast
                done
            qed
          qed
        have "{?l} Un ?S1 :netlists"
          apply(rule addGate)
          apply(rule ccontr)
          apply(simp)
          apply(subgoal_tac "?S1 ⊆  nl")
          apply(subgoal_tac "isDefinedIn n nl")
          apply(cut_tac b1, simp)
          apply(unfold isDefinedIn_def,simp add:subsetNetUn[THEN sym]) thm subsetNetRef
          apply(blast intro: subsetNetRef)+
          apply(assumption)
          apply(cut_tac d6,fold isDefinedIn_def,assumption)
          apply(cut_tac d5,assumption)
          done
 
        with d4  have "{?l} Un ?S2 :netlists" by auto
        with d2 d3  have ?thesis by auto
      }
      moreover 
      {assume 
        d1:"?l~: subNetList ({?l} ∪ nl) snl"
        have "subNetList ({?l}Un nl) (snl)=subNetList ( nl) (snl)"
          apply(rule subsetNet2)
          apply(cut_tac d1)
          by assumption
        also have "…=subNetList ( nl) (snl Int nl)"
          by (rule  subsetNetJoint1)
        also have "subNetList ( nl) (snl Int nl):netlists"
          apply(cut_tac b3)
          apply(drule_tac x="?S0" in spec)
          apply blast
          done
        ultimately have ?thesis by simp
      }
      ultimately show ?thesis by blast
    qed
  qed
qed

lemma subNetIsNetlist:
  assumes
  a1:"nl:netlists" and
  a2:"snl ⊆ nl" 
  shows  
  "   (subNetList nl snl):netlists"
  (is  "?P nl")
  apply -
  apply(cut_tac a1,  drule_tac nl="nl" in  subNetIsNetlist)
  apply(cut_tac a2,blast)
  done

end
     

lemma subsetNetRef:

  x ∈ subNetList nl nl' ==> xnl

lemma subsetNetlistNil:

  x ∈ subNetList {} snl ==> False

lemma subsetNetUn1:

  x ∈ subNetList nl (snl1.0snl2.0) ==>
  x ∈ subNetList nl snl1.0x ∈ subNetList nl snl2.0

lemma subsetUnRev1:

  x ∈ subNetList nl snl1.0x ∈ subNetList nl snl2.0 ==>
  x ∈ subNetList nl (snl1.0snl2.0)

lemma subsetNetUn:

  subNetList nl (snl1.0snl2.0) = subNetList nl snl1.0 ∪ subNetList nl snl2.0

lemma subsetNetAdd:

  x ∈ subNetList ({l} ∪ nl) snl ==>
  x ∈ subNetList nl snlx = lx ∈ subNetList nl {l0nl. fanout l0 ∈ set (fanins l)}

lemma subsetNetAux:

  [|x ∈ subNetList nl {l0nl. fanout l0 ∈ set (fanins l)};
   l ∈ subNetList ({l} ∪ nl) snl|]
  ==> x ∈ subNetList ({l} ∪ nl) snl

lemma subsetNetMono:

  [|x ∈ subNetList M nl; MN|] ==> x ∈ subNetList N nl

lemma subsetNetMono2:

  [|x ∈ subNetList M nl; nlnl'|] ==> x ∈ subNetList M nl'

lemma subsetNet1:

  l ∈ subNetList ({l} ∪ nl) snl ==>
  subNetList ({l} ∪ nl) snl =
  subNetList nl snl ∪ {l} ∪ subNetList nl {l0nl. fanout l0 ∈ set (fanins l)}

lemma subsetNetAux1:

  [|l ∉ subNetList ({l} ∪ nl) snl; x ∈ subNetList ({l} ∪ nl) snl|]
  ==> x ∈ subNetList nl snl

lemma subsetNet2:

  l ∉ subNetList ({l} ∪ nl) snl ==> subNetList ({l} ∪ nl) snl = subNetList nl snl

lemma subsetNetlistNil:

  subNetList {} snl = {}

lemma subsetNetJoint:

  x ∈ subNetList nl nl' ==> x ∈ subNetList nl (nl'nl)

lemma subsetNetJoint1:

  subNetList nl nl' = subNetList nl (nl'nl)

lemma subNetIsNetlist:

  nl ∈ netlists ==> ∀snlnl. subNetList nl snl ∈ netlists

lemma subNetIsNetlist:

  [|nl ∈ netlists; snlnl|] ==> subNetList nl snl ∈ netlists