Theory caseCam

Up to index of Isabelle/HOL/net

theory caseCam
imports caseParaComparator
begin

theory caseCam=caseParaComparator:

   

lemma posAssAndSet:
  
  shows " 
   ALL M N i.  len = M - N --> 
  set (PosAssertOfLine (map match [N ..<M])
  ((λ i. (map (λ j. (if (i=j) then ONE else DontCare)) [N ..<M])) ( i))) ⊆ 
  {chaos} Un {Is1 (match (i::nat))}" (is "?P len")
proof(induct_tac len)
  show "?P 0"
  proof(simp del:cktSat_def)qed
  fix n
  assume b0:"?P n"
  show "?P (Suc n)"
  proof((rule allI)+,(rule impI)+)
    fix M N i
    assume  b1:"Suc n = M - N" 
    have b3:"n=M - (Suc N)" by(cut_tac b1,arith)
    have b4:"EX M'.  M=Suc M'& N <=M'"
      by(cut_tac b1,arith)
    then obtain M' where b5:"M=Suc M'& N <=M'"
      by blast
    then  have b6:"[N..<M]=[ N..<M']@[M']" 
      apply - 
      apply auto
      done
     have b7:"[N..<M]=N#[Suc N ..<M]"
     proof(rule upt_conv_Cons,cut_tac b5,simp)qed
     show "set (PosAssertOfLine (map match [N..<M])
                     (map (λj. if i = j then ONE else DontCare) [N..<M]))
           ⊆ {chaos} ∪ {Is1 (match i)}"
       apply(cut_tac b0 b7,case_tac "i = N",simp add:Let_def,
         drule_tac x="M" in spec,
         drule_tac x="Suc N"  in spec,
         cut_tac b3,simp,simp add:Let_def,
         drule_tac x="M" in spec,
         drule_tac x="Suc N"  in spec,
         cut_tac b3,simp)
       done
   qed
qed

lemma posAssAndSet:
  shows " 
  set (PosAssertOfLine (map match [N ..<M])
  ((λ i. (map (λ j. (if (i=j) then ONE else DontCare)) [N ..<M])) ( i))) ⊆ 
  {chaos} Un {Is1 (match (i::nat))}"
proof( cut_tac posAssAndSet[where len="M - N" and match="match"] , 
    drule_tac x="M" in spec ,
    drule_tac x="N" in spec,
    drule_tac x="i"in spec, simp )qed
    
        
lemma camAux1:
  assumes a1:"nl:netlists" and
  a2:"isClosed nl" and
  a3:"orLine=(λ i. (map (λ j. (if (i=j) then ONE else DontCare)) [0..<M]))" and
  a5:"matches=map match [0..<M]"
 shows "cktSat nl ((Is1 (match i)) \<leadsto>
  andLists (PosAssertOfLine matches (orLine i)))" (is "?P i")
proof(rule_tac B="andLists [Is1 (match i)]" in steTrans)
  show "cktSat nl (Is1 (match i) \<leadsto> andLists [Is1 (match i)])" thm andListsAllI1
    proof(rule andListsAllI)
      show "∀a. a ∈ set [Is1 (match i)] --> cktSat nl (Is1 (match i) \<leadsto> a)"
      proof(force intro:steRef)qed
    qed
next
  show "cktSat nl (andLists [Is1 (match i)] \<leadsto> andLists (PosAssertOfLine matches (orLine i)))"
  proof(rule andListSubsetEq)
    have "set (PosAssertOfLine (map match [0 ..<M])
      ((λ i. (map (λ j. (if (i=j) then ONE else DontCare)) [0 ..<M])) ( i))) ⊆ 
      {chaos} Un {Is1 (match (i::nat))}"
    proof(rule  posAssAndSet)qed
    then show "set (PosAssertOfLine matches (orLine i)) ∩ - {chaos} ⊆ set [Is1 (match i)]"
      apply -
      apply(cut_tac a3 a5,simp,force)
      done
    qed
  qed
   

lemma camCase:
  assumes a0:"css=map (λi. (map (λj. c i j) [0..<N])) [0..<M]" and
  a0a:"xorGs={g. EX j i. j<N &i<M &g=Gate (c i j) [(Tag j),(storedTag i j) ] xorTab}" and
  a1:"xorTab=[[ONE,ONE],[ZERO,ZERO]]" and a2:"andTab=[andLine]" and
  a3:"isClosed nl" and
  a6:"andLine=map (λ n. ONE) [0..<N]" and
  a7:"andGs={g. EX j i .j<N &i<M & g=Gate (match i)  (map (λj. c i j) [0..<N]) andTab}"   and 
  a8:" (xorGs Un andGs Un {orG}  )⊆ nl  " and
  a9:"nl:netlists" and
  a10:"orLine=(λ i. (map (λ j. (if (i=j) then ONE else DontCare)) [0..<M]))" and
  a11:"orTab=map orLine [0..<M]" and
  a12:"matches= map match [0..<M]" and
  a13:"orG=Gate hit matches orTab" and
  a14:"antOfTag = map (λj. Isb (Tag j) (bvOfTag j)) [0..<N]" and
  a15:"antOfTs = map (λi. (map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]))  [0..<M]" and
  a16:"GpOfUnHitI=(λ i. (EX j. j<N & (bvOfTag j)~=( bvOfTs i j)))" and
  a17:"GpOfHitI=(λ i. (ALL j. j<N --> (bvOfTag j)=( bvOfTs i j)))" and
  a18:"ant=andLists (antOfTag @ (concat antOfTs))" and
  a19:"cons0=GpOfUnHit ->T (Is0 hit)" and
  a20:"cons1=GpOfHit  ->T (Is1 hit)" and
  a21:"GpOfUnHit=(ALL i. i<M --> (GpOfUnHitI i))" and
  a22:"GpOfHit=(EX i. i<M & (GpOfHitI i))" and
  a23:"0<M" and
  a24:"0<N"
  shows "cktSat nl (ant \<leadsto> (cons0 andT cons1) )" 
proof(rule steConjI)
  let ?asLs="map (Is0) matches"
  show " cktSat nl (ant \<leadsto> cons0)" 
  proof(cut_tac a19 a19,simp del:cktSat_def,rule steImpI1)
    assume b0:"GpOfUnHit"
    show "cktSat nl (ant \<leadsto> Is0 hit)" thm orTabPropF
    proof(rule_tac ?B="(andLists ?asLs)" in steTrans)
      show " cktSat nl (andLists (map Is0 matches) \<leadsto> Is0 hit)"
      proof(rule_tac ?asLists=" ?asLs" in orTabPropF)
        show "Gate hit matches orTab ∈ nl" 
          by(cut_tac a13 a8, simp)
        show "∀l. l ∈ set orTab --> length l = length matches"
        proof(cut_tac a12 a11 a10,simp,force)qed
        show "∀line. line mem orTab --> (∃as. as mem map Is0 matches ∧ isNegAssertOfLine as matches line)"
        proof(rule allI,rule impI)
          fix line
          assume c1:"line mem orTab"
          show " ∃as. as mem map Is0 matches ∧ isNegAssertOfLine as matches line"
          proof -
            from c1 have "EX i. i<M & line =(orLine i)"
            proof(cut_tac a10 a11 a12 c1, simp add:mem_iff,force)qed
            then obtain i where d1:"i<M & line =(orLine i)"
              by blast
            show " ∃as. as mem map Is0 matches ∧ isNegAssertOfLine as matches line" thm exI
            proof(rule_tac x="Is0 (match i)" in  exI)
              show "Is0 (match i) mem map Is0 matches ∧ 
                isNegAssertOfLine (Is0 (match i)) matches line"
              proof(cut_tac a10 a11 a12 d1,simp add:mem_iff set_zip,force)qed
            qed
          qed
        qed
      qed
      show "cktSat nl (ant \<leadsto> andLists (map Is0 matches))"
      proof(rule andListsAllI)
        show  a1:"∀a. a ∈ set (map Is0 matches) --> cktSat nl (ant \<leadsto> a)"
        proof(rule allI,rule impI)
          fix a
          assume b1:" a ∈ set (map Is0 matches)"
          from b1 have "EX i. i<M & a=Is0 (match i)"
          proof(cut_tac a10 a11 a12 b1,simp add:mem_iff,force)qed
          then obtain i where d1:"i<M &a=Is0 (match i)"
            by blast
          let ?antOfAs="antOfTag"
          let ?antOfBs=" (map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N])"
          let ?ant="andLists (?antOfAs @?antOfBs)"
          show "cktSat nl (ant \<leadsto> a)"
          proof(cut_tac d1,simp del:cktSat_def,
              rule_tac ?B="?ant" in steTrans)
            show " cktSat nl (ant \<leadsto> andLists
              (antOfTag @ map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]))"
            proof(cut_tac a18 a15,simp del:cktSat_def,rule andListCat   )
              show "cktSat nl ((andLists  antOfTag) \<leadsto> andLists antOfTag)"
                by(rule steRef)
              show "cktSat nl
                (andLists
                (concat (map (λi. map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]) [0..<M])) \<leadsto>
                andLists (map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]))"
                proof(rule andListSubsetEq)
                  show " set (map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N])∩ - {chaos}
                    ⊆ set (concat (map (λi. map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]) [0..<M]))"
                  proof(cut_tac d1,simp del:cktSat_def,force)qed
                qed
              qed
              let ?xorGs="{g. EX j . j<N &i<M &g=Gate (c i j) [(Tag j),(storedTag i j)] xorTab}"
              let ?c="c i" 
              let ?n="N" 
              let ?cs="map (c i) [0..<N]" 
              let ?andG="Gate (match i) ?cs andTab"
              let ?b="storedTag i" 
              let ?a="Tag"
              let ?bvOfAs="bvOfTag "
              let ?bvOfBs="bvOfTs i"
              let ?antOfAs =" map (λj. Isb (Tag j) (bvOfTag j)) [0..<N]"
              let ?antOfBs ="map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]"
              let ?Gp0.0 = "(∃j<N. bvOfTag j ≠bvOfTs i j )"
              let ?Gp1.0 =" (∀j<N.  bvOfTag j= bvOfTs i j)"
              let ?cons0.0 = "(∃j<N. bvOfTag j ≠ (bvOfTs i j)) ->T Is0 (match i)"
              let ?cons1.0 = "(∀j<N. bvOfTag j = (bvOfTs i j)) ->T Is1 (match i)"
              show "cktSat nl (andLists (antOfTag @ map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]) \<leadsto> Is0 (match i))"      
              proof(rule_tac  ?xorTab ="xorTab" and ?andLine="andLine" and
                  ?andTab="andTab" and ?xorGLs="?xorGs" and ?c="?c" and
                  ?n="N" and ?cs="?cs" and ?andG="?andG" and ?a="?a" and
                  ?b="?b" and ?bvOfAs="?bvOfAs" and ?bvOfBs="?bvOfBs" and
                  ?antOfAs="?antOfAs" and ?antOfBs ="?antOfBs" and
                  ?Gp0.0="?Gp0.0" and ?Gp1.0 ="?Gp1.0" and?cons0.0="?cons0.0" and
                  ?cons1.0 = "?cons1.0"
                  in caseParaComparatorNeq,(assumption)+,cut_tac d1,simp,
                  cut_tac a6,simp,(force)+, cut_tac d1 a8 a0a a7 a24 a23,simp,force,(assumption)+,(simp+),
                  (assumption)+,simp,
                  cut_tac a14,simp
                )
                show "∃j<N. bvOfTag j ≠ bvOfTs i j"
                proof(cut_tac d1 b0 a16 a21,simp)qed
              qed
            qed
          qed
        qed
      qed
    qed
    show "cktSat nl (ant \<leadsto> cons1)"
    proof(cut_tac a20 a20,simp del:cktSat_def,rule steImpI1)
      assume b0:"GpOfHit"
      from b0 obtain i where b1:"i<M &GpOfHitI i"
      proof(cut_tac a22,simp,auto)qed
      show "cktSat nl (ant \<leadsto> Is1 hit)" 
      proof(rule_tac ?B="andLists (PosAssertOfLine matches (orLine i))" in steTrans)
        show " cktSat nl (ant \<leadsto> andLists (PosAssertOfLine matches (orLine i)))"
        proof(rule_tac ?B="Is1 (match i)" in steTrans)
          show "cktSat nl (Is1 (match i) \<leadsto>andLists ( PosAssertOfLine matches (orLine i)))" 
          proof(rule camAux1) qed
        next
           let ?xorGs="{g. EX j . j<N &i<M &g=Gate (c i j) [(Tag j),(storedTag i j)] xorTab}"
              let ?c="c i" 
              let ?n="N" 
              let ?cs="map (c i) [0..<N]" 
              let ?andG="Gate (match i) ?cs andTab"
              let ?b="storedTag i" 
              let ?a="Tag"
              let ?bvOfAs="bvOfTag "
              let ?bvOfBs="bvOfTs i"
              let ?antOfAs =" map (λj. Isb (Tag j) (bvOfTag j)) [0..<N]"
              let ?antOfBs ="map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]"
              let ?Gp0.0 = "(∃j<N. bvOfTag j ≠bvOfTs i j )"
              let ?Gp1.0 =" (∀j<N.  bvOfTag j= bvOfTs i j)"
              let ?cons0.0 = "(∃j<N. bvOfTag j ≠ (bvOfTs i j)) ->T Is0 (match i)"
              let ?cons1.0 = "(∀j<N. bvOfTag j = (bvOfTs i j)) ->T Is1 (match i)"
              let ?ant="andLists (?antOfAs@?antOfBs)"
              have c1:"cktSat nl (?ant \<leadsto> Is1 (match i))" thm caseParaComparatorEq      
              proof(rule_tac  ?xorTab ="xorTab" and ?andLine="andLine" and
                  ?andTab="andTab" and ?xorGLs="?xorGs" and ?c="?c" and
                  ?n="N" and ?cs="?cs" and ?andG="?andG" and ?a="?a" and
                  ?b="?b" and ?bvOfAs="?bvOfAs" and ?bvOfBs="?bvOfBs" and
                  ?antOfAs="?antOfAs" and ?antOfBs ="?antOfBs" and
                  ?Gp0.0="?Gp0.0" and ?Gp1.0 ="?Gp1.0" and?cons0.0="?cons0.0" and
                  ?cons1.0 = "?cons1.0" 
                  in caseParaComparatorEq,(assumption)+,
                  cut_tac b1,simp,
                  cut_tac a6,simp,(force)+
                  , cut_tac b1 a8 a0a a7 a24 a23,simp,
                  force,
                  (assumption)+,(simp+),
                  (assumption)+,simp+
                )
                show " ∀j<N. bvOfTag j = bvOfTs i j"
                proof(cut_tac b0 b1 a17,force)qed
              qed
             have c2:"cktSat nl (ant \<leadsto>  ?ant)"
             proof(cut_tac a18 a15,simp del:cktSat_def,rule andListCat  )
              show "cktSat nl ((andLists  antOfTag) \<leadsto> andLists (map (λj. Isb (Tag j) (bvOfTag j)) [0..<N]))"
                by(cut_tac a14, simp del:cktSat_def, rule steRef)
              show "cktSat nl
                (andLists
                (concat (map (λi. map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]) [0..<M])) \<leadsto>
                andLists (map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]))"
                proof(rule andListSubsetEq)
                  show " set (map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N])∩ - {chaos}
                    ⊆ set (concat (map (λi. map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]) [0..<M]))"
                  proof(cut_tac b1,simp del:cktSat_def,force)qed
                qed
              qed
             with c1 show " cktSat nl (ant \<leadsto> Is1 (match i))"
             proof(rule_tac steTrans)qed
           qed
         next
           show "cktSat nl (andLists (PosAssertOfLine matches (orLine i)) \<leadsto> Is1 hit)"
           proof(rule orTabPropT)
             show "Gate hit matches orTab ∈ nl" 
               by(cut_tac a13 a8, simp)
             show "∀l. l ∈ set orTab --> length l = length matches"
             proof(cut_tac a12 a11 a10,simp,force)qed 
             show "orLine i mem orTab"
             proof(cut_tac a10 a11 b1,simp add:mem_iff,force)qed
           qed
         qed
       qed 
     qed
end

lemma posAssAndSet:

M N i.
     len = M - N -->
     set (PosAssertOfLine (map match [N..<M])
           (map (λj. if i = j then ONE else DontCare) [N..<M]))
     ⊆ {chaos} ∪ {Is1 (match i)}

lemma posAssAndSet:

  set (PosAssertOfLine (map match [N..<M])
        (map (λj. if i = j then ONE else DontCare) [N..<M]))
  ⊆ {chaos} ∪ {Is1 (match i)}

lemma camAux1:

  [|nl ∈ netlists; isClosed nl;
   orLine = (λi. map (λj. if i = j then ONE else DontCare) [0..<M]);
   matches = map match [0..<M]|]
  ==> cktSat nl
       (Is1 (match i) \<leadsto> andLists (PosAssertOfLine matches (orLine i)))

lemma camCase:

  [|css = map (λi. map (c i) [0..<N]) [0..<M];
   xorGs =
   {g. ∃j i. j < Ni < Mg = Gate (c i j) [Tag j, storedTag i j] xorTab};
   xorTab = [[ONE, ONE], [ZERO, ZERO]]; andTab = [andLine]; isClosed nl;
   andLine = map (λn. ONE) [0..<N];
   andGs =
   {g. ∃j i. j < Ni < Mg = Gate (match i) (map (c i) [0..<N]) andTab};
   xorGsandGs ∪ {orG} ⊆ nl; nl ∈ netlists;
   orLine = (λi. map (λj. if i = j then ONE else DontCare) [0..<M]);
   orTab = map orLine [0..<M]; matches = map match [0..<M];
   orG = Gate hit matches orTab;
   antOfTag = map (λj. Isb (Tag j) (bvOfTag j)) [0..<N];
   antOfTs = map (λi. map (λj. Isb (storedTag i j) (bvOfTs i j)) [0..<N]) [0..<M];
   GpOfUnHitI = (λi. ∃j<N. bvOfTag jbvOfTs i j);
   GpOfHitI = (λi. ∀j<N. bvOfTag j = bvOfTs i j);
   ant = andLists (antOfTag @ concat antOfTs); cons0.0 = GpOfUnHit ->T Is0 hit;
   cons1.0 = GpOfHit ->T Is1 hit; GpOfUnHit = (∀i<M. GpOfUnHitI i);
   GpOfHit = (∃i<M. GpOfHitI i); 0 < M; 0 < N|]
  ==> cktSat nl (ant \<leadsto> cons0.0 andT cons1.0)