Theory valuePart

Up to index of Isabelle/HOL/net

theory valuePart
imports Main
begin

theory valuePart=Main:
types boolPairs="bool × bool"
constdefs Top::"boolPairs" 
         "Top==(False,False)"
constdefs T::"boolPairs"
         "T==(True,False)"
constdefs F::"boolPairs"
         "F==(False,True)"
constdefs X::"boolPairs"
         "X==(True,True)"
constdefs lub::"(boolPairs) =>  (boolPairs) => (boolPairs)" 
"lub a b ==(fst a&fst b,snd a&snd b)"

constdefs OR::"boolPairs =>boolPairs =>boolPairs"
"OR x y==(( (fst x ) | (fst y)),((snd x ) & snd y))    "     
constdefs AND::"boolPairs =>boolPairs =>boolPairs"
"AND x y==(( (fst x ) & (fst y)),((snd x ) | snd y))    "   
constdefs NOT::"boolPairs =>boolPairs"
"NOT x==((snd x), fst x)"

types node=nat
types time=nat
types state=" node=> boolPairs"

types stateSeq= "nat=>state "




constdefs stateLub::"state =>state =>state"
"stateLub s1 s2==λ n. (lub (s1 n) (s2 n)) "


constdefs stateSeqLub::"stateSeq=>stateSeq=>stateSeq"
"stateSeqLub seq1 seq2 ==
  λ t. stateLub (seq1 t) (seq2 t)"

constdefs leq::"boolPairs => boolPairs =>bool" (infix "\<sqsubseteq> " 80)
"leq a b==(b=lub a b)"


constdefs stateLeq::"state =>state =>bool" (infix "\<sqsubseteq>s" 80)
"stateLeq s1 s2 == ALL n.( (s1 n)\<sqsubseteq>  (s2 n))"

constdefs stateSqLeq::"stateSeq =>stateSeq =>bool" (infix "\<sqsubseteq>sq " 80)
"stateSqLeq sq1 sq2==ALL t. (sq1) t \<sqsubseteq>s  (sq2) t"

constdefs suffix::"time => stateSeq => stateSeq"
"suffix i sq== λ t.( sq (t+i) )"


constdefs drop ::"bool => boolPairs"
"drop x==if (x=True) then T else F"

constdefs stateDrop ::"  (node => bool) =>state "
"stateDrop s== λ n. drop (s n)"


constdefs stateSqDrop ::"  (time=>node => bool) =>stateSeq"
"stateSqDrop sq== λ t. stateDrop (sq t)"



datatype LIT=ONE|ZERO|DontCare
(*types literal="LIT× node"*)
types LINE="LIT list"
types  PLA="LINE list"

constdefs funOfLit::" boolPairs × LIT=>boolPairs"
 "funOfLit x == if (snd x)=ONE then (fst x)
                        else if (snd x)=ZERO then (NOT (fst x))
                        else T"

consts funOfLine::" boolPairs list => LINE =>boolPairs"
primrec 
  "funOfLine  bps []=T"
  "funOfLine  (bps) (el0#ls)= AND (funOfLit ((hd bps),el0)) (funOfLine  (tl bps) (ls))
  "
  
consts funOfTab::"PLA => boolPairs list => boolPairs"
primrec
  "funOfTab [] bps= F"
  "funOfTab (l#pla) bps=OR (funOfLine bps l) (funOfTab (pla) bps)"


constdefs isFullAndLine::"LINE =>bool"
  "isFullAndLine line0 == ALL l. l mem line0 --> l=ONE"


constdefs isAndTab::"PLA =>bool"
  "isAndTab tab == length tab=1 & isFullAndLine (hd tab)"




consts PosValOfLine::" LINE  =>boolPairs list"
primrec
PosVal1:  " PosValOfLine  [] =([])"

PosVal2:  " PosValOfLine  (lit#line)=
   (let otherAss=PosValOfLine  ( line) in
   (case (lit) of ONE => (T)#otherAss|
                     ZERO  => (F)#otherAss|
                     DontCare  => X#otherAss))"




 


lemma tl_map:"(tl (map s1 inps)) =map s1 (tl inps)"
apply (induct_tac inps)
apply simp
apply simp
done

lemma andMono:"[|a1 \<sqsubseteq> a2; b1\<sqsubseteq> b2|] ==> (AND a1 b1)\<sqsubseteq> (AND a2 b2)"
apply -
apply(unfold AND_def leq_def)
apply(unfold lub_def)
apply(simp add:Pair_fst_snd_eq)
apply auto 
done

lemma andF:"F \<sqsubseteq> v ==>  F\<sqsubseteq>(AND  v  b)" 
apply(unfold leq_def lub_def AND_def F_def,simp add:Pair_fst_snd_eq)

done

lemma andT: "[|T \<sqsubseteq> v1; T \<sqsubseteq> v2|] ==>  T\<sqsubseteq>(AND  v1  v2)"
apply(unfold leq_def lub_def AND_def T_def,simp add:Pair_fst_snd_eq)
done
lemma andComm:"AND a b=AND b a" 
apply(unfold leq_def lub_def AND_def F_def,simp add:Pair_fst_snd_eq)
apply blast
done

lemma orMono:"[|a1 \<sqsubseteq> a2; b1\<sqsubseteq> b2|] ==> (OR a1 b1)\<sqsubseteq> (OR a2 b2)"
apply -
apply(unfold OR_def leq_def)
apply(unfold lub_def)
apply(simp add:Pair_fst_snd_eq)
apply auto 
done

lemma orT:"[|T \<sqsubseteq> v1|]==>T \<sqsubseteq> (OR  v1 v2)"
apply (unfold OR_def leq_def)
apply(unfold lub_def T_def)
apply(simp add:Pair_fst_snd_eq)
done

lemma orF:"[|F \<sqsubseteq> v1;F \<sqsubseteq> v2|]==>  F\<sqsubseteq> (OR v1 v2)"
apply (unfold OR_def leq_def)
apply(unfold lub_def F_def)
apply(simp add:Pair_fst_snd_eq)
done

lemma orComm:"OR a b=OR b a" 
apply(unfold leq_def lub_def OR_def F_def,simp add:Pair_fst_snd_eq)
apply blast
done
lemma notMono:"[|a1 \<sqsubseteq> a2|] ==> (NOT a1 )\<sqsubseteq> (NOT a2 )"
  apply(unfold NOT_def leq_def)
  apply(unfold lub_def)
  apply(simp add:Pair_fst_snd_eq)  
  done

lemma lubMono:"[|a1 \<sqsubseteq> a2; b1\<sqsubseteq> b2|] ==> (lub a1 b1)\<sqsubseteq> (lub a2 b2)"
  apply(unfold NOT_def leq_def)
  apply(unfold lub_def)
  apply(simp add:Pair_fst_snd_eq)  
  apply auto
  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 lubSelf: " a \<sqsubseteq> (lub a  b )"
  apply(unfold leq_def lub_def)
  apply auto
done

lemma lubId:" lub a a=a"
  by(unfold  lub_def,  auto)

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



lemma XIsLeastValue:
  " X \<sqsubseteq> a"
  apply(unfold leq_def)
  apply(unfold lub_def)
  apply(unfold X_def,simp)
done

lemma lubMeans:
  "[| x \<sqsubseteq> a; y\<sqsubseteq> a|] ==> lub x y  \<sqsubseteq> a"
  apply(unfold leq_def lub_def)
  apply auto 
  apply(simp add:Pair_fst_snd_eq)
  done
  

lemma leqReflexive:"a \<sqsubseteq> a"
  apply(unfold leq_def)
  apply(unfold lub_def)
  apply(simp add:Pair_fst_snd_eq)  
  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 sqLeqIsTrans:
"[|sq0  \<sqsubseteq>sq sq1; sq1  \<sqsubseteq>sq sq2 |]==> sq0  \<sqsubseteq>sq sq2"
  apply(unfold stateSqLeq_def stateLeq_def)
  apply(rule allI)+
  apply(blast intro:leqIsTrans)
done


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

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

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


lemma sqLeqIsReflexive:" sq \<sqsubseteq>sq sq"
  apply(unfold stateSqLeq_def  stateLeq_def)
  apply(rule allI)+
  apply(rule leqReflexive)
  done


lemma funOfLitIsMono:"a1 \<sqsubseteq>a2 ==>funOfLit(a1,sign0) \<sqsubseteq>funOfLit(a2,sign0)"
  apply(case_tac sign0)
  apply(unfold funOfLit_def)
  apply simp
  apply simp
  apply(simp add:notMono) 
  apply(simp )
  apply auto
  apply(simp add:leqReflexive)
  done

lemma hdMap1:"inps~=[]-->hd (map s1 inps)=s1 (hd inps) & (hd inps) mem inps"
  apply(induct_tac inps)
  apply simp
  apply simp
done

lemma hdMap:"inps~=[]==>hd (map s1 inps)=s1 (hd inps) & (hd inps) mem inps"
  apply(cut_tac hdMap1[where inps="inps" and ?s1.0="s1" ])
  apply simp
 
done

lemma funOfLineIsMono:"(s1 \<sqsubseteq>s s2) -->
  ( ALL inps. length inps=length line-->
   (funOfLine (map s1 inps)) line\<sqsubseteq>  (funOfLine (map s2 inps)) line )"
apply(induct_tac line)
apply simp
apply(unfold leq_def lub_def)
apply simp
apply(fold  lub_def leq_def)

apply simp
apply(rule impI)+
apply(rule allI,rule impI)
apply(subgoal_tac "(tl (map s1 inps)) =map s1 (tl inps) ")
apply(subgoal_tac "(tl (map s2 inps)) =map s2 (tl inps)")
prefer 2
apply(rule tl_map)
apply simp
prefer 2
apply(rule tl_map) thm  andMono
apply(rule andMono)
apply(subgoal_tac "hd (map s1 inps)=s1 (hd inps)& (hd inps) mem inps")
apply(subgoal_tac "hd (map s2 inps)=s2 (hd inps)& (hd inps) mem inps")
apply(rule funOfLitIsMono)
apply(unfold stateLeq_def)
apply simp
apply(rule hdMap)
apply(subgoal_tac "0<length inps")
apply(simp (no_asm_use)  add: length_greater_0_conv )
apply simp
apply(rule hdMap)
apply(subgoal_tac "0<length inps")
apply(simp (no_asm_use)  add: length_greater_0_conv )
apply simp
apply auto
done

lemma funOfLineIsMono:"[|(s1 \<sqsubseteq>s s2); length inps=length tab|]==>  
   (funOfLine (map s1 inps)) tab\<sqsubseteq>  (funOfLine (map s2 inps)) tab "
apply(cut_tac funOfLineIsMono[where ?s1.0="s1" and ?s2.0="s2" and line="tab"])
apply blast
done

lemma funOfLineIsMonoAny:"
  ( ALL inps. length inps=length line-->(ALL n. n mem inps --> s1 n \<sqsubseteq> s2 n) -->
   (funOfLine (map s1 inps)) line\<sqsubseteq>  (funOfLine (map s2 inps)) line )"
apply(induct_tac line)
apply simp
apply(unfold leq_def lub_def)
apply simp
apply(fold  lub_def leq_def)

apply simp
apply(rule allI,(rule impI)+)
apply(subgoal_tac "(tl (map s1 inps)) =map s1 (tl inps)")
apply(subgoal_tac "(tl (map s2 inps)) =map s2 (tl inps)")
prefer 2
apply(rule tl_map)
apply simp
prefer 2
apply(rule tl_map) thm  andMono
apply(rule andMono)
apply(subgoal_tac "hd (map s1 inps)=s1 (hd inps)& (hd inps) mem inps")
apply(subgoal_tac "hd (map s2 inps)=s2 (hd inps)& (hd inps) mem inps")
apply(rule funOfLitIsMono)
apply simp
apply(rule hdMap)
apply(subgoal_tac "0<length inps")
apply(simp (no_asm_use)  add: length_greater_0_conv )
apply simp
apply(rule hdMap)
apply(subgoal_tac "0<length inps")
apply(simp (no_asm_use)  add: length_greater_0_conv )
apply simp
apply(drule_tac x="(tl inps)" in spec)
apply simp
apply(subgoal_tac "(∀n. n mem tl inps --> s1 n \<sqsubseteq>  s2 n)")
apply blast
apply(rule allI)
apply(rule impI)
apply(subgoal_tac "n mem inps")
apply blast
apply(case_tac inps)
apply auto
done

lemma funOfLineIsMonoAny:"[|(ALL n. n mem inps --> s1 n \<sqsubseteq> s2 n); length inps=length tab|]==>  
   (funOfLine (map s1 inps)) tab\<sqsubseteq>  (funOfLine (map s2 inps)) tab "
apply(cut_tac funOfLineIsMonoAny[where ?s1.0="s1" and ?s2.0="s2" and line="tab"])
apply blast
done

lemma funOfTabIsMono:
shows "(s1 \<sqsubseteq>s s2) -->(ALL l. (l:set tab)--> length inps=length l) -->
  funOfTab tab (map s1 inps) \<sqsubseteq>  funOfTab tab (map s2 inps)"
apply -
apply(induct_tac tab)
apply simp
apply((rule impI)+,rule leqReflexive)
apply simp
apply((rule impI)+,rule orMono)
apply(rule funOfLineIsMono)
apply blast+
done

lemma funOfTabIsMono:
shows "[|(s1 \<sqsubseteq>s s2) ;(ALL l. (l:set tab)--> length inps=length l) |]==>
  funOfTab tab (map s1 inps) \<sqsubseteq>  funOfTab tab (map s2 inps)"
apply(cut_tac  funOfTabIsMono[where ?s1.0="s1" and ?s2.0="s2" and tab="tab"])
apply blast
done

lemma funOfTabIsMonoAny:
shows "(ALL n. n mem inps --> s1 n \<sqsubseteq> s2 n) -->(ALL l. (l:set tab)--> length inps=length l) -->
  funOfTab tab (map s1 inps) \<sqsubseteq>  funOfTab tab (map s2 inps)"
apply -
apply(induct_tac tab)
apply simp
apply((rule impI)+,rule leqReflexive)
apply simp
apply((rule impI)+,rule orMono)
apply(rule funOfLineIsMonoAny)
apply blast+
done

lemma funOfTabIsMonoAny:
shows "[|(ALL n. n mem inps --> s1 n \<sqsubseteq> s2 n);(ALL l. (l:set tab)--> length inps=length l) |]==>
  funOfTab tab (map s1 inps) \<sqsubseteq>  funOfTab tab (map s2 inps)"
apply(cut_tac  funOfTabIsMonoAny[where ?s1.0="s1" and ?s2.0="s2" and tab="tab"])
apply blast
done

consts stateSqLeq1::"boolPairs list => boolPairs list  => bool"
primrec
stateSqLeq1Nil: "stateSqLeq1 [] ls= True"
stateSqLeq1Cons:"stateSqLeq1 (x#xs) ys = ((x \<sqsubseteq> (hd ys)) & (stateSqLeq1 xs (tl ys)))"

lemma hdStateSqLeq1:
"[|stateSqLeq1 stateLs1 stateLs2; length stateLs1 = Suc (length list)|]
       ==> hd stateLs1 \<sqsubseteq>  hd stateLs2"
apply(simp (no_asm_use)  only:length_Suc_conv )
apply(erule exE)+
apply(simp add:stateSqLeq1Cons)
done

lemma funOfLineIsMono:
"ALL stateLs1 stateLs2. stateSqLeq1 stateLs1   stateLs2 -->
  ( length  stateLs1=length l-->(funOfLine stateLs1 l) \<sqsubseteq>  (funOfLine stateLs2 l))"
apply(induct_tac l)
apply simp
apply(unfold leq_def lub_def)
apply simp
apply(fold  lub_def leq_def)

apply simp
apply((rule allI)+,(rule impI)+)
apply(rule andMono)
apply(rule funOfLitIsMono)
apply(rule hdStateSqLeq1)
apply assumption+ 
apply(drule_tac x="tl stateLs1" in spec)
apply(drule_tac x="tl stateLs2" in spec)
apply(simp (no_asm_use)  only:length_Suc_conv )
apply(erule exE)+
apply(simp add:stateSqLeq1Cons)
done

lemma funOfLineIsMono:
"[| stateSqLeq1 stateLs1   stateLs2 ;
  length  stateLs1=length l|] ==>((funOfLine stateLs1 l) \<sqsubseteq>  (funOfLine stateLs2 l))"
apply(cut_tac funOfLineIsMono[where l="l"])
apply blast
done

lemma andLinePropF:
assumes a:"F \<sqsubseteq> v"
shows 
"ALL bps.  v mem bps --> length bps=length line --> isFullAndLine line-->F \<sqsubseteq> funOfLine bps line"
  (is "?P line")
proof(induct_tac line)
  show "?P []"
    apply(rule allI,rule impI,rule impI)
    apply simp
    done
next 
  fix a list
  assume a0:"?P list"
  show "?P (a # list)"
    proof(rule allI,(rule impI)+)
      fix bps
      assume a1:"v mem bps" and
      a2:"length bps = length (a # list)" and
      a3:"isFullAndLine (a # list)"
      show "F \<sqsubseteq>funOfLine bps (a # list)"
      proof(simp)
    let ?P=
     "  F \<sqsubseteq>AND (funOfLit (hd bps, a)) (funOfLine (tl bps) list) "
    from a2 have b1:"EX n. length bps=Suc n"
      apply auto
      done
    then obtain n where b1:"length bps=Suc n"
      by auto
    then have b2:"EX y ys.( bps = y # ys & length ys = n)"
      by (simp add:length_Suc_conv)
    then  obtain y ys  where b2:"bps = y # ys "
      by auto
    from a1 b2 have "y=v | v mem ys"
      apply -
      apply(case_tac "y=v", auto)
      done
    moreover
    {assume b3:"y=v"
      have b4:"a=ONE"
        apply(cut_tac a3,unfold isFullAndLine_def)
        apply auto
        done
      have " F \<sqsubseteq>(funOfLit (hd bps, a))"
        apply(cut_tac  b3 b2 b4 ,simp add:funOfLit_def)
        apply(cut_tac a,simp)
        done
      then have ?P
        apply - 
        apply(rule_tac andF) 
        apply (simp)
        done
    }
    moreover
    {assume b3:" v mem ys"
      have b4: "length ys = length list "
        apply(cut_tac a2 b2)
        by auto

      have "isFullAndLine list "
        apply(cut_tac a3)
        apply(unfold isFullAndLine_def)
        by auto
      
      from b3 a0 this b4  have "F \<sqsubseteq> funOfLine ys list "
        by auto
      from b2 this have "F \<sqsubseteq>funOfLine (tl bps) list "
        by auto
      then have ?P
        apply -
        apply(subgoal_tac "AND (funOfLit (hd bps, a)) (funOfLine (tl bps) list)=
          AND (funOfLine (tl bps) list) (funOfLit (hd bps, a))")
        apply simp
        apply(rule_tac andF) 
        apply (simp)
        apply (rule andComm)
        done
      }
      ultimately show ?P by blast
    qed
      qed
    qed

lemma andLinePropF:
"[|F \<sqsubseteq> v; v mem bps ; length bps=length line; isFullAndLine line|]==> F \<sqsubseteq> funOfLine bps line"
  apply(cut_tac  andLinePropF[where line="line"])
  apply blast
  apply simp
  done



(*a n usual line, need not be a fullAndLine*)
lemma aLinePropT:
shows 
"ALL bps.   length bps=length line --> bps= PosValOfLine line-->T\<sqsubseteq> funOfLine bps line"
  (is "?P line")
proof(induct_tac line)
  show "?P []"
    apply(rule allI,rule impI,rule impI)
    apply simp
    apply(rule leqReflexive)
    done
next 
  fix a list
  assume a0:"?P list"
  show "?P ((a::LIT) # list)"
    proof(rule allI,(rule impI)+)
      fix bps
      assume a2:"length bps = length (a # list)" and
        a3:"bps = PosValOfLine  (a # list)" 
      show "T \<sqsubseteq> funOfLine bps (a # list) "
      proof(simp)
        let ?P=
          " T \<sqsubseteq> AND (funOfLit (hd bps, a)) (funOfLine (tl bps) list) "
        from a2 have b1:"EX n. length bps=Suc n"
          apply auto
          done
        then obtain n where b1:"length bps=Suc n"
          by auto
        then have b2:"EX y ys.( bps = y # ys & length ys = n)"
          by (simp add:length_Suc_conv)
        then  obtain y ys  where b2:"bps = y # ys "
          by auto
        have b3:" T \<sqsubseteq>(funOfLit (hd bps, (a))) &  ys = PosValOfLine ( list)"
          apply(cut_tac a3 b2  ,simp add:funOfLit_def)   
          apply(case_tac "a::LIT", simp add:Let_def)
          apply(rule leqReflexive)
          apply( simp add:Let_def)
          apply(unfold T_def F_def leq_def lub_def NOT_def,simp add:Pair_fst_snd_eq)
          apply (simp add:Let_def)
          done
        have b4a: "length ys = length list "
          apply(cut_tac a2 b2)
          by auto

        
      
        from b3 a0 this b4a  have "T \<sqsubseteq>funOfLine ys list "
          by auto
        from b2 this have "T \<sqsubseteq>funOfLine (tl bps) list"
          by auto
        from b3 this show ?P
          apply -
          by (rule andT,auto)
      
      qed
    qed
  qed

lemma aLinePropT:
shows 
"[| length bps=length line ; bps= PosValOfLine line|]==>T\<sqsubseteq> funOfLine bps line"
proof(cut_tac  aLinePropT[where line="line"],blast) qed

consts listLeq::"boolPairs list => boolPairs list => bool"
primrec 
  listLeq1:"listLeq [] l2=True"
  listLeq2:"listLeq (a#list0) l =(case l of []=> False |
                            (b#list1)=> (a \<sqsubseteq> b)& (listLeq list0 list1))"
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 listLeqTrans:" ALL list2 list3. listLeq list1 list2 --> listLeq list2 list3-->listLeq list1 list3"
  apply(induct_tac list1)
  apply simp
  apply(rule allI)
  apply(induct_tac list2)
  apply simp
  apply(rule allI)
  apply(induct_tac list3)
  apply simp
  apply(rule impI)+
  apply( simp (no_asm))
  apply(simp)
  apply(erule conjE)+
  apply(rule conjI)
  apply(rule_tac b="aa" in  leqIsTrans)
  apply simp+
  done

lemma listLeqTrans:"
[|listLeq list1 list2 ; listLeq list2 list3|]==>listLeq list1 list3"
apply(cut_tac  listLeqTrans[where ?list1.0="list1"])
apply blast
done







lemma aLinePropT1:
shows 
"ALL bps.   length bps=length line --> listLeq (PosValOfLine line) bps-->T\<sqsubseteq> funOfLine bps line"
  (is "?P line")
proof(induct_tac line)
  show "?P []"
    apply(rule allI,rule impI,rule impI)
    apply simp
    apply(rule leqReflexive)
    done
next 
  fix a list
  assume a0:"?P list"
  show "?P ((a::LIT) # list)"
    proof(rule allI,(rule impI)+)
      fix bps
      assume a2:"length bps = length (a # list)" and
        a3:"listLeq (PosValOfLine (a # list)) bps" 
      show "T \<sqsubseteq> funOfLine bps (a # list) "
      proof(simp)
        let ?P=
          " T \<sqsubseteq> AND (funOfLit (hd bps, a)) (funOfLine (tl bps) list) "
        from a2 have b1:"EX n. length bps=Suc n"
          apply auto
          done
        then obtain n where b1:"length bps=Suc n"
          by auto
        then have b2:"EX y ys.( bps = y # ys & length ys = n)"
          by (simp add:length_Suc_conv)
        then  obtain y ys  where b2:"bps = y # ys "
          by auto
        have b3:" T \<sqsubseteq>(funOfLit (hd bps, (a))) & listLeq (PosValOfLine list) ys"
          apply(cut_tac a3 b2  ,simp add:funOfLit_def)   
          apply(case_tac "a::LIT", simp add:Let_def)
          apply( simp add:Let_def)
          apply(unfold T_def F_def leq_def lub_def NOT_def,simp add:Pair_fst_snd_eq)
          apply (simp add:Let_def)
          done
        have b4a: "length ys = length list "
          apply(cut_tac a2 b2)
          by auto

        
      
        from b3 a0 this b4a  have "T \<sqsubseteq>funOfLine ys list "
          by auto
        from b2 this have "T \<sqsubseteq>funOfLine (tl bps) list"
          by auto
        from b3 this show ?P
          apply -
          by (rule andT,auto)
      
      qed
    qed
  qed



lemma aLinePropT1:
shows 
"[| length bps=length line ;listLeq (PosValOfLine line) bps|]==>T\<sqsubseteq> funOfLine bps line"
proof(cut_tac  aLinePropT1[where line="line"],blast) qed



lemma andLinePropT: 
  
  "ALL bps.  (ALL bv. bv mem bps --> T \<sqsubseteq>bv ) -->
  length bps=length line --> isFullAndLine line-->  T \<sqsubseteq> funOfLine bps line"
  (is "?P line")
proof(induct_tac line)
  show "?P []"
    apply(rule allI,rule impI,rule impI)
    apply simp
    apply(rule impI, rule leqReflexive)
    done
next 
  fix a list
  assume a0:"?P list"
  show "?P (a # list)"
    proof(rule allI,(rule impI)+)
      fix bps
      assume a1:"(∀bv. bv mem bps --> T \<sqsubseteq>bv )" and
      a2:"length bps = length (a # list)" and
      a3:"isFullAndLine (a # list)"
      show "T \<sqsubseteq> funOfLine bps (a # list) "
      proof(simp)
    let ?P=
     " T \<sqsubseteq> AND (funOfLit (hd bps, a)) (funOfLine (tl bps) list) "
    from a2 have b1:"EX n. length bps=Suc n"
      apply auto
      done
    then obtain n where b1:"length bps=Suc n"
      by auto
    then have b2:"EX y ys.( bps = y # ys & length ys = n)"
      by (simp add:length_Suc_conv)
    then  obtain y ys  where b2:"bps = y # ys "
      by auto
    from a1 b2 have b3:"T \<sqsubseteq>y &( ALL y. y mem ys -->T \<sqsubseteq>y)"
      apply -
      apply (simp del:pair_collapse split_paired_All)     
      done
    

    have b4:"a=ONE"
      apply(cut_tac a3,unfold isFullAndLine_def)
      apply auto
      done
    have b5:"T \<sqsubseteq>(funOfLit (hd bps, a))"
      apply(cut_tac b3 b2 b4 ,simp add:funOfLit_def)
      done
    
    have b4a: "length ys = length list "
      apply(cut_tac a2 b2)
      by auto

    have b4b:"isFullAndLine list "
      apply(cut_tac a3)
      apply(unfold isFullAndLine_def)
      by auto
      
    from b3 a0 this b4a  have "T \<sqsubseteq>funOfLine ys list "
      by auto
    from b2 this have "T \<sqsubseteq>funOfLine (tl bps) list"
        by auto
    from b5 this show ?P
      apply -
      by (rule andT)
      
      qed
    qed
  qed

lemma andLinePropT:
  "[|(ALL bv. bv mem bps --> T \<sqsubseteq>bv) ;
  length bps=length line; isFullAndLine line|]==> T \<sqsubseteq> funOfLine bps line"
  apply(cut_tac  andLinePropT[where line="line"])
  apply blast
  done

lemma andTabPropF:
"[|F \<sqsubseteq>v; v mem bps ; length bps=length (hd tab) ; isAndTab tab|]==>F \<sqsubseteq> funOfTab tab bps"
apply(unfold isAndTab_def)
apply simp
apply(erule conjE)
apply(subgoal_tac "EX y ys.( tab = y # ys & length ys = 0)")
apply(erule exE)+
apply(erule conjE)
apply simp
apply(subgoal_tac "F \<sqsubseteq>funOfLine bps y")
apply(simp add:OR_def F_def)
apply(rule andLinePropF)
apply assumption+
apply(simp add:length_Suc_conv)
done
lemma orT:"[|T \<sqsubseteq>a|] ==>T \<sqsubseteq> OR  a b "
apply(unfold leq_def lub_def OR_def T_def,simp add:Pair_fst_snd_eq)
done
lemma andTabPropT:
"[|ALL bv. bv mem bps -->  T \<sqsubseteq>bv ; length bps=length (hd tab) ; isAndTab tab|]==>  T \<sqsubseteq>funOfTab tab bps "
apply(unfold isAndTab_def)
apply( simp del:pair_collapse split_paired_All)
apply(erule conjE)
apply(subgoal_tac "EX y ys.( tab = y # ys & length ys = 0)")
apply(erule exE)+
apply(erule conjE)
apply(simp del:pair_collapse split_paired_All)
apply(subgoal_tac " T \<sqsubseteq>funOfLine bps y")
apply(rule orT)
apply assumption
apply(rule andLinePropT)
apply assumption+
apply(simp add:length_Suc_conv)
done

lemma orTabPropT:
shows "ALL bps. line mem tab-->length bps=length line -->
  bps= PosValOfLine line--> T\<sqsubseteq> funOfTab tab bps " (is "?P tab")
proof(induct_tac tab)
  show "?P []"
  proof(simp) qed
next
  fix a list
  assume a0:"?P list"
  show "?P (a # list)"
  proof(rule allI, (rule impI)+)
    fix bps
    assume a1:"line mem a # list" and
      a2:"length bps = length line" and
      a3:" bps = PosValOfLine line"
    show "T \<sqsubseteq> funOfTab (a # list) bps"
    proof -
      have "a=line | line mem list" 
        apply(cut_tac a1,simp)
        by(case_tac "a=line",auto)
      moreover
      {assume b1:"a=line"
        have b2:"T\<sqsubseteq> funOfLine bps line"
        proof(rule       aLinePropT) qed
        have "T \<sqsubseteq>  funOfTab (a # list) bps"
        proof(simp,rule_tac orT,cut_tac b1 b2,auto) qed
      }
      moreover
      {assume b1:"line mem list"
        from a0 have b2:" T\<sqsubseteq> funOfTab list bps "
          apply -
          apply(drule_tac x="bps" in spec)
          apply(cut_tac a2 a3 b1)
          by blast
        then have "T \<sqsubseteq>  funOfTab (a # list) bps"
        proof(simp) 
          have c1:"OR (funOfLine bps a) (funOfTab list bps)=
            OR  (funOfTab list bps) (funOfLine bps a)"
            by (rule_tac orComm)
          have c2:"T \<sqsubseteq> OR  (funOfTab list bps) (funOfLine bps a)"
          proof(rule_tac orT,cut_tac  b2,auto) qed
          with c1 show "T \<sqsubseteq>OR (funOfLine bps a) (funOfTab list bps)"
            by auto
        qed
      }
      ultimately   show "T \<sqsubseteq> funOfTab (a # list) bps"
        by blast
    qed
  qed
qed
      

lemma orTabPropT:
 "[| line mem tab;length bps=length line ;
  bps= PosValOfLine line|] ==>  T\<sqsubseteq> funOfTab tab bps "
apply(cut_tac   orTabPropT[where tab="tab" and line="line"],blast)
done
 lemma orTabPropT1:
shows "ALL bps. line mem tab-->length bps=length line -->
   listLeq (PosValOfLine line) bps--> T\<sqsubseteq> funOfTab tab bps " (is "?P tab")
proof(induct_tac tab)
  show "?P []"
  proof(simp) qed
next
  fix a list
  assume a0:"?P list"
  show "?P (a # list)"
  proof(rule allI, (rule impI)+)
    fix bps
    assume a1:"line mem a # list" and
      a2:"length bps = length line" and
      a3:" listLeq (PosValOfLine line) bps"
    show "T \<sqsubseteq> funOfTab (a # list) bps"
    proof -
      have "a=line | line mem list" 
        apply(cut_tac a1,simp)
        by(case_tac "a=line",auto)
      moreover
      {assume b1:"a=line"
        have b2:"T\<sqsubseteq> funOfLine bps line"
        proof(rule       aLinePropT1) qed
        have "T \<sqsubseteq>  funOfTab (a # list) bps"
        proof(simp,rule_tac orT,cut_tac b1 b2,auto) qed
      }
      moreover
      {assume b1:"line mem list"
        from a0 have b2:" T\<sqsubseteq> funOfTab list bps "
          apply -
          apply(drule_tac x="bps" in spec)
          apply(cut_tac a2 a3 b1)
          by blast
        then have "T \<sqsubseteq>  funOfTab (a # list) bps"
        proof(simp) 
          have c1:"OR (funOfLine bps a) (funOfTab list bps)=
            OR  (funOfTab list bps) (funOfLine bps a)"
            by (rule_tac orComm)
          have c2:"T \<sqsubseteq> OR  (funOfTab list bps) (funOfLine bps a)"
          proof(rule_tac orT,cut_tac  b2,auto) qed
          with c1 show "T \<sqsubseteq>OR (funOfLine bps a) (funOfTab list bps)"
            by auto
        qed
      }
      ultimately   show "T \<sqsubseteq> funOfTab (a # list) bps"
        by blast
    qed
  qed
qed
      

lemma orTabPropT:
 "[| line mem tab;length bps=length line ;
 listLeq (PosValOfLine line) bps|] ==>  T\<sqsubseteq> funOfTab tab bps "
apply(cut_tac   orTabPropT1[where tab="tab" and line="line"],blast)
done

types LITValuePair="LIT×  boolPairs"

constdefs contradict:: "LITValuePair => bool"
"contradict pair ≡  (fst pair)=ONE &   F\<sqsubseteq>(snd pair) |
 (fst pair)=ZERO & T\<sqsubseteq>(snd pair)"

lemma aLinePropF:
shows
"ALL bps.   length bps=length line --> 
  (EX asOfVal. asOfVal mem ( zip line bps) & contradict asOfVal)-->F \<sqsubseteq> funOfLine bps line"
  (is "?P line")
proof(induct_tac line)
  show "?P []" by auto
next
  fix a line
  assume b1:"?P line"
  show "?P (a#line)"
  proof(rule allI, (rule impI)+)
    fix bps
    assume c1:"length bps = length (a # line)" and
    c2:"∃asOfVal. asOfVal mem zip (a # line) bps ∧ contradict asOfVal" 
    show " F \<sqsubseteq>  funOfLine bps (a # line)"
    proof -
      have " length bps =Suc (length line) "
        proof(cut_tac c1, auto) qed
        then have d2:"EX y ys.( bps = y # ys & length ys = length line)"
          by (simp add:length_Suc_conv)
        then  obtain z zs  where d2:"bps = z # zs& length zs = length line"
          by auto

      from c2 obtain asOfVal where 
        d1:"asOfVal mem zip (a # line) bps ∧ contradict asOfVal" 
        by blast
      
      with d2 have "asOfVal=(a,z)| asOfVal mem zip ( line) zs" 
        apply -
        apply simp
        by(case_tac "(a, z) = asOfVal",auto)
      moreover
      {assume e1:"asOfVal=(a,z)"
        with d1 d2 have ?thesis
        proof(simp add:contradict_def funOfLit_def,case_tac "a::LIT",auto,
              rule_tac andF,simp, rule_tac andF,
              simp add:F_def T_def NOT_def leq_def lub_def,
              simp add:Pair_fst_snd_eq) 
        qed
      } 
      moreover
      {assume e1:"asOfVal mem zip line zs"
        with d1 d2 b1 have e2: "F \<sqsubseteq> funOfLine zs line"
          by(drule_tac x="zs" in spec,blast)
        with d2  have ?thesis
        proof(simp)
          have " AND (funOfLit (z, a)) (funOfLine zs line)=
             AND  (funOfLine zs line)(funOfLit (z, a))"
            by (rule andComm)
          with  e2 show   "F \<sqsubseteq>AND (funOfLit (z, a)) (funOfLine zs line)"
            apply - apply(simp)
            apply(rule_tac andF)
            by simp
        qed
      }
      ultimately show ?thesis by blast
      qed
    qed
  qed

lemma aLinePropF:
  assumes a1:"length bps=length line" and
  a2:"(EX asOfVal. asOfVal mem ( zip line bps) & contradict asOfVal)"
  shows
  "F \<sqsubseteq> funOfLine bps line"
proof(cut_tac aLinePropF[where line="line"] a1 a2,blast) qed

lemma orTabPropF1:
   
  shows "ALL bps. ( (∀ l. (l∈ (set tab)) --> length l=length bps)-->
  (ALL line. line mem  tab --> (EX asOfVal. asOfVal mem ( zip line bps) & contradict asOfVal))  -->
  F \<sqsubseteq> funOfTab tab bps)" (is "?P tab")
proof(induct_tac tab)
  show "?P []" by ( simp,rule leqReflexive)
next
  fix a tab
  assume b1:"?P tab"
  show "?P (a#tab)"
  proof(rule allI,(rule impI)+)
    fix bps
    assume c1:"∀l. l ∈ set (a # tab) --> length l = length bps" and
    c2:"∀line. line mem a # tab --> (∃asOfVal. asOfVal mem zip line (bps:: boolPairs list)
      ∧ contradict asOfVal)"
    show " F \<sqsubseteq>  funOfTab (a # tab) bps"
    proof(simp)
      from c2 have c3:"(∃asOfVal. asOfVal mem zip a (bps:: boolPairs list)
      ∧ contradict asOfVal)"
      proof(drule_tac x="a" in spec,simp) qed
      from c2 
      have c4:"∀line. line mem  tab --> (∃asOfVal. asOfVal mem zip line (bps:: boolPairs list)
        ∧ contradict asOfVal)"
      proof(rule_tac allI,rule_tac impI,force) qed
      from c1 have c5:"length a=length bps" by simp
      from c1 have c6:"∀ l. l∈ (set tab) --> length l=length bps"
        by simp
      from c3 c5 have c7:"F \<sqsubseteq>(funOfLine bps a)" 
        by(rule_tac  aLinePropF,auto)
      from c4 c6 b1 have "F \<sqsubseteq>  funOfTab ( tab) bps"
      proof(drule_tac x="bps" in spec, blast)qed
      with c7 show " F \<sqsubseteq>  OR (funOfLine bps a) (funOfTab tab bps)"
      proof(rule_tac orF)qed
    qed
  qed
qed
        
lemma orTabPropF1:
  assumes a1:" (∀ l. (l∈ (set tab)) --> length l=length bps)" and
  a2:"(ALL line. line mem  tab --> (EX asOfVal. asOfVal mem ( zip line bps) & contradict asOfVal))"

  shows "  F \<sqsubseteq> funOfTab tab bps"
proof(cut_tac orTabPropF1[where tab="tab"] a1 a2, blast)qed

lemma arith1:" 0< x ==> EX n. x= Suc n" by arith

consts consListAux::"'a list => node list"
primrec 
consAuxNil:" consListAux []=[]"
consAuxCons:"consListAux (x#xs)=(length xs +1) # (consListAux xs)"

lemma conAuxSameLen:" length (consListAux xs)=length xs"
  apply(induct_tac xs)
  apply auto
done


end 

lemma tl_map:

  tl (map s1.0 inps) = map s1.0 (tl inps)

lemma andMono:

  [|a1.0 \<sqsubseteq>  a2.0; b1.0 \<sqsubseteq>  b2.0|]
  ==> AND a1.0 b1.0 \<sqsubseteq>  AND a2.0 b2.0

lemma andF:

  F \<sqsubseteq>  v ==> F \<sqsubseteq>  AND v b

lemma andT:

  [|T \<sqsubseteq>  v1.0; T \<sqsubseteq>  v2.0|]
  ==> T \<sqsubseteq>  AND v1.0 v2.0

lemma andComm:

  AND a b = AND b a

lemma orMono:

  [|a1.0 \<sqsubseteq>  a2.0; b1.0 \<sqsubseteq>  b2.0|]
  ==> OR a1.0 b1.0 \<sqsubseteq>  OR a2.0 b2.0

lemma orT:

  T \<sqsubseteq>  v1.0 ==> T \<sqsubseteq>  OR v1.0 v2.0

lemma orF:

  [|F \<sqsubseteq>  v1.0; F \<sqsubseteq>  v2.0|]
  ==> F \<sqsubseteq>  OR v1.0 v2.0

lemma orComm:

  OR a b = OR b a

lemma notMono:

  a1.0 \<sqsubseteq>  a2.0 ==> NOT a1.0 \<sqsubseteq>  NOT a2.0

lemma lubMono:

  [|a1.0 \<sqsubseteq>  a2.0; b1.0 \<sqsubseteq>  b2.0|]
  ==> lub a1.0 b1.0 \<sqsubseteq>  lub a2.0 b2.0

lemma lubSelf:

  a \<sqsubseteq>  lub a b

lemma lubComm:

  lub a b = lub b a

lemma lubSelf:

  a \<sqsubseteq>  lub a b

lemma lubId:

  lub a a = a

lemma lubComm:

  lub a b = lub b a

lemma XIsLeastValue:

  X \<sqsubseteq>  a

lemma lubMeans:

  [|x \<sqsubseteq>  a; y \<sqsubseteq>  a|] ==> lub x y \<sqsubseteq>  a

lemma leqReflexive:

  a \<sqsubseteq>  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 sqLeqIsTrans:

  [|sq0.0 \<sqsubseteq>sq  sq1.0; sq1.0 \<sqsubseteq>sq  sq2.0|]
  ==> sq0.0 \<sqsubseteq>sq  sq2.0

lemma eqIByLeq:

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

lemma Leq1I:

  a = b ==> a \<sqsubseteq>  b

lemma Leq2I:

  a = b ==> b \<sqsubseteq>  a

lemma sqLeqIsReflexive:

  sq \<sqsubseteq>sq  sq

lemma funOfLitIsMono:

  a1.0 \<sqsubseteq>  a2.0 ==>
  funOfLit (a1.0, sign0.0) \<sqsubseteq>  funOfLit (a2.0, sign0.0)

lemma hdMap1:

  inps ≠ [] --> hd (map s1.0 inps) = s1.0 (hd inps) ∧ hd inps mem inps

lemma hdMap:

  inps ≠ [] ==> hd (map s1.0 inps) = s1.0 (hd inps) ∧ hd inps mem inps

lemma funOfLineIsMono:

  s1.0 \<sqsubseteq>s s2.0 -->
  (∀inps. length inps = length line -->
          funOfLine (map s1.0 inps) line \<sqsubseteq> 
          funOfLine (map s2.0 inps) line)

lemma funOfLineIsMono:

  [|s1.0 \<sqsubseteq>s s2.0; length inps = length tab|]
  ==> funOfLine (map s1.0 inps) tab \<sqsubseteq>  funOfLine (map s2.0 inps) tab

lemma funOfLineIsMonoAny:

inps. length inps = length line -->
         (∀n. n mem inps --> s1.0 n \<sqsubseteq>  s2.0 n) -->
         funOfLine (map s1.0 inps) line \<sqsubseteq> 
         funOfLine (map s2.0 inps) line

lemma funOfLineIsMonoAny:

  [|∀n. n mem inps --> s1.0 n \<sqsubseteq>  s2.0 n; length inps = length tab|]
  ==> funOfLine (map s1.0 inps) tab \<sqsubseteq>  funOfLine (map s2.0 inps) tab

lemma funOfTabIsMono:

  s1.0 \<sqsubseteq>s s2.0 -->
  (∀l. l ∈ set tab --> length inps = length l) -->
  funOfTab tab (map s1.0 inps) \<sqsubseteq>  funOfTab tab (map s2.0 inps)

lemma funOfTabIsMono:

  [|s1.0 \<sqsubseteq>s s2.0; ∀l. l ∈ set tab --> length inps = length l|]
  ==> funOfTab tab (map s1.0 inps) \<sqsubseteq>  funOfTab tab (map s2.0 inps)

lemma funOfTabIsMonoAny:

  (∀n. n mem inps --> s1.0 n \<sqsubseteq>  s2.0 n) -->
  (∀l. l ∈ set tab --> length inps = length l) -->
  funOfTab tab (map s1.0 inps) \<sqsubseteq>  funOfTab tab (map s2.0 inps)

lemma funOfTabIsMonoAny:

  [|∀n. n mem inps --> s1.0 n \<sqsubseteq>  s2.0 n;
   ∀l. l ∈ set tab --> length inps = length l|]
  ==> funOfTab tab (map s1.0 inps) \<sqsubseteq>  funOfTab tab (map s2.0 inps)

lemma hdStateSqLeq1:

  [|stateSqLeq1 stateLs1.0 stateLs2.0; length stateLs1.0 = Suc (length list)|]
  ==> hd stateLs1.0 \<sqsubseteq>  hd stateLs2.0

lemma funOfLineIsMono:

stateLs1 stateLs2.
     stateSqLeq1 stateLs1 stateLs2 -->
     length stateLs1 = length l -->
     funOfLine stateLs1 l \<sqsubseteq>  funOfLine stateLs2 l

lemma funOfLineIsMono:

  [|stateSqLeq1 stateLs1.0 stateLs2.0; length stateLs1.0 = length l|]
  ==> funOfLine stateLs1.0 l \<sqsubseteq>  funOfLine stateLs2.0 l

lemma andLinePropF:

  F \<sqsubseteq>  v ==>
  ∀bps. v mem bps -->
        length bps = length line -->
        isFullAndLine line --> F \<sqsubseteq>  funOfLine bps line

lemma andLinePropF:

  [|F \<sqsubseteq>  v; v mem bps; length bps = length line; isFullAndLine line|]
  ==> F \<sqsubseteq>  funOfLine bps line

lemma aLinePropT:

bps. length bps = length line -->
        bps = PosValOfLine line --> T \<sqsubseteq>  funOfLine bps line

lemma aLinePropT:

  [|length bps = length line; bps = PosValOfLine line|]
  ==> T \<sqsubseteq>  funOfLine bps line

lemma leqIsTrans:

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

lemma listLeqTrans:

list2 list3.
     listLeq list1.0 list2 --> listLeq list2 list3 --> listLeq list1.0 list3

lemma listLeqTrans:

  [|listLeq list1.0 list2.0; listLeq list2.0 list3.0|] ==> listLeq list1.0 list3.0

lemma aLinePropT1:

bps. length bps = length line -->
        listLeq (PosValOfLine line) bps --> T \<sqsubseteq>  funOfLine bps line

lemma aLinePropT1:

  [|length bps = length line; listLeq (PosValOfLine line) bps|]
  ==> T \<sqsubseteq>  funOfLine bps line

lemma andLinePropT:

bps. (∀bv. bv mem bps --> T \<sqsubseteq>  bv) -->
        length bps = length line -->
        isFullAndLine line --> T \<sqsubseteq>  funOfLine bps line

lemma andLinePropT:

  [|∀bv. bv mem bps --> T \<sqsubseteq>  bv; length bps = length line;
   isFullAndLine line|]
  ==> T \<sqsubseteq>  funOfLine bps line

lemma andTabPropF:

  [|F \<sqsubseteq>  v; v mem bps; length bps = length (hd tab); isAndTab tab|]
  ==> F \<sqsubseteq>  funOfTab tab bps

lemma orT:

  T \<sqsubseteq>  a ==> T \<sqsubseteq>  OR a b

lemma andTabPropT:

  [|∀bv. bv mem bps --> T \<sqsubseteq>  bv; length bps = length (hd tab);
   isAndTab tab|]
  ==> T \<sqsubseteq>  funOfTab tab bps

lemma orTabPropT:

bps. line mem tab -->
        length bps = length line -->
        bps = PosValOfLine line --> T \<sqsubseteq>  funOfTab tab bps

lemma orTabPropT:

  [|line mem tab; length bps = length line; bps = PosValOfLine line|]
  ==> T \<sqsubseteq>  funOfTab tab bps

lemma orTabPropT1:

bps. line mem tab -->
        length bps = length line -->
        listLeq (PosValOfLine line) bps --> T \<sqsubseteq>  funOfTab tab bps

lemma orTabPropT:

  [|line mem tab; length bps = length line; listLeq (PosValOfLine line) bps|]
  ==> T \<sqsubseteq>  funOfTab tab bps

lemma aLinePropF:

bps. length bps = length line -->
        (∃asOfVal. asOfVal mem zip line bps ∧ contradict asOfVal) -->
        F \<sqsubseteq>  funOfLine bps line

lemma aLinePropF:

  [|length bps = length line;
   ∃asOfVal. asOfVal mem zip line bps ∧ contradict asOfVal|]
  ==> F \<sqsubseteq>  funOfLine bps line

lemma orTabPropF1:

bps. (∀l. l ∈ set tab --> length l = length bps) -->
        (∀line. line mem tab -->
                (∃asOfVal. asOfVal mem zip line bps ∧ contradict asOfVal)) -->
        F \<sqsubseteq>  funOfTab tab bps

lemma orTabPropF1:

  [|∀l. l ∈ set tab --> length l = length bps;
   ∀line. line mem tab -->
          (∃asOfVal. asOfVal mem zip line bps ∧ contradict asOfVal)|]
  ==> F \<sqsubseteq>  funOfTab tab bps

lemma arith1:

  0 < x ==> ∃n. x = Suc n

lemma conAuxSameLen:

  length (consListAux xs) = length xs