Theory nsl05

Up to index of Isabelle/HOL/strand

theory nsl05
imports strand05 Public
begin

theory nsl05= strand05+ Public:
constdefs is_initiator::"sigma =>  agent=> agent =>nat =>nat =>bool" 
  "is_initiator s A B Na Nb == 
  (SP s)=[(+, Crypt (pubK B) {|Nonce Na, Agent A|}), 
  (-,Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|}),
  (+, Crypt (pubK B) (Nonce Nb))]"

constdefs is_responder::"sigma =>  agent=> agent =>nat =>nat =>bool" 
  "is_responder s A B Na Nb ==
  (SP s)=[(-, Crypt (pubK B) {|Nonce Na, Agent A|}),
  (+,Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|}),
  (-, Crypt (pubK B) (Nonce Nb))]"

defs NSL_1:
"Σ == {s. Is_penetrator_strand s |
          (EX A B Na Nb.  is_initiator s  A B Na Nb) |
          (EX A B Na Nb.  is_responder s A B Na Nb)
       }"



lemma unique_orig_init:
  "[|is_initiator s A B Na Nb ; is_initiator s' A' B' Na Nb';
  unique_originate (Nonce Na) (s,0)|]
  ==>s=s'& A=A'& B=B'&Nb=Nb'"
  apply(unfold is_initiator_def
    unique_originate_def) 
  apply(subgoal_tac "originate (Nonce Na) (s',0)")
  apply(erule conjE)+
  apply(drule_tac x="(s',0)" in spec)
  apply simp
  apply(unfold originate_def)
  apply(unfold node_sign_def node_term_def
    parts_relation_def index_def,auto)
  done

lemma no_node_is_key: 
  assumes a1:"b:bundles" and a2:"n: nodes b"  
  shows c:"ALL k. ( ~k: KP) -->  ~node_term n=Key k"
proof (rule_tac P="ALL k. (( ~k: KP) -->  ~node_term n=Key k)" in  ccontr)
  assume a3:"¬ (∀k. ( ~k:KP) --> node_term n ≠ Key k) " 
  show "False"
  proof -
    from a3 have M0:"EX k. ( ~k:KP )&node_term n = Key k" 
      by auto
    
    from this obtain k 
      where b1: "( ~k:KP) & node_term n = Key k " 
      by auto
    let ?P="{x. Key k \<sqsubset> node_term x & x:nodes b &( ~k:KP)}" 
    from b1 and a2   have M2:"n: ?P"
      by( simp, blast intro:subterm_itself)
    from  a1 and this 
    have M1:"EX z:?P . (ALL y.(y,z):(edges b)^+ -->y~:?P)"  (is "EX z:?P. ?Q z") 
      apply - apply (drule_tac x="n" and  Q="?P" in bundle_minimal)
      apply simp 
      apply blast
      done
    
    from this obtain m'
      where M2: "m':?P &?Q m'"
      by auto
    from this have M3:" node_sign m'= +"
    proof(rule_tac P="node_sign m' = +" in  ccontr)
      assume a00:"node_sign m' ≠ +" 
      from a00 have M4:"node_sign m'=-" 
        apply - 
        apply (cut_tac node_sign_is_Plus_or_minus[where n="node_sign m'"],
          blast) done

      from a1 and M2 and M4
      have M5:"EX m''. m'':nodes b & m'' ->  m'&(m'', m') ∈ edges b" 
        apply - apply ((erule conjE)+,simp only:mem_Collect_eq) 
        apply( drule_tac b="b" and ?n2.0="m'" in bundle_edge_is_exist)
        apply blast done 
      from this obtain m'' where M6:"m'':nodes b & m'' ->  m'&(m'', m') ∈ edges b"
        by auto
      from this have M7:"node_term m'=node_term m''" 
        apply - apply (unfold casual1_def,auto)
        done
      from this and M2 and M6   
      have " m'':?P" by (simp only:mem_Collect_eq, blast)
 
      from this and M2 and M6 show  "False"  
        apply - thm r_into_trancl 
        apply(erule conjE)+ 
        apply(drule_tac x="m''" in spec) 
        apply (blast dest:r_into_trancl)
        done
    qed
    from a1 and M2 have a8:"m':Domain" 
      by(blast dest: bundle_node_in_Domain)
    let ?s="strand  m'"
    have "(Is_penetrator_strand ?s | ~Is_penetrator_strand ?s)" 
      by auto
    moreover
    {assume a4:" Is_penetrator_strand ?s"
      from a4 have "(  Is_Tee_strand ?s | Is_Flush_strand ?s | Is_Cat_strand ?s |Is_Sep_strand ?s|
        Is_E_strand ?s | Is_D_strand ?s |Is_T_strand ?s | Is_K_strand ?s)"
        apply - apply(unfold  penetrator_trace) 
        apply auto done
      moreover
      { assume a5:"Is_E_strand (?s)"
        from a5  and M3 and a8  have M0:"snd m'=2" 
          apply - 
          apply( unfold Is_E_strand_def, unfold Domain_def,
            unfold node_sign_def index_def strand_def,auto)
          apply(case_tac "y=0")
          apply simp 
          apply (case_tac "y=1" ) 
          apply simp 
          apply auto  
          done
        from a5 and this 
        have M1:"EX k h . node_term (fst  m',0)=Key k &
                          node_term (fst m', 1)=h &  
                          node_term m'=Crypt k h" (is "EX k h. ?P k h")
          by (unfold node_term_def, 
              unfold Is_E_strand_def _strand_def
              index_def, 
              auto)
        then obtain k' and h where M01:" ?P k' h" 
          by auto
        
        from M0  and a8 and M2 and a1 
        have M000:"(fst m',1):fst b&((fst m', 1) ,m'):(snd b)^*"
          apply-  apply(simp)
          apply((erule conjE)+) 
          apply(subgoal_tac "1<= snd m'") 
          apply(fold strand_def index_def nodes_def edges_def)
          apply( drule_tac i="1" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
          apply simp
          apply simp  
          apply(erule conjE, drule_tac x="((strand m', 1), m')"
            and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
          apply blast+
          apply simp+  
          done

        from this and M0 have M3:"((fst m',1),m'): (snd b )^+" 
          apply -  
          apply(erule conjE,drule rtranclD) 
          apply(erule disjE)
          apply(fold strand_def index_def)
          apply( simp add:node_equal) 
          apply(unfold strand_def index_def)
          apply auto
          done

        from M01 and M2 
        have M4:"Key k \<sqsubset> Crypt k' h"
          by auto

        from this have M5:"Key k \<sqsubset>  h" 
          apply - apply(drule body_subterm, auto)
          done

        from this and M2 and M3 and M01 and M000 
        have ?thesis 
          apply- apply(simp only:mem_Collect_eq)
          apply(erule conjE)+
          apply(fold edges_def nodes_def)
          apply(drule_tac x="(fst m', 1)" in spec)
          apply  auto done
      }
      moreover
      { assume a5:"Is_D_strand ?s"
        from a5 and a8 and M3  
        have M0:"snd m'=2" 
          apply - 
          apply( unfold Is_D_strand_def, 
            unfold Domain_def, unfold node_sign_def
            strand_def index_def ,auto)
          apply(case_tac "y=0") 
          apply simp   
          apply (case_tac "y=1" ) 
          apply simp 
          apply auto   done
        from a5 and this
        have M1:"EX k. EX k'. EX h. k'=invKey k& 
          node_term (fst m',0)= (Key k')&node_term (fst m',1)= (Crypt k h)&
          node_term  m'=h" (is "EX k k'  h. ?P k k' h")
          by (unfold node_term_def, 
            unfold Is_D_strand_def 
            strand_def index_def, 
            auto)
        then obtain k' and k'' and h 
          where M01:" ?P k'  k'' h" 
          by auto
        from a5 
        have M00:"node_sign (fst m', 0)=-& node_sign (fst m', 1)=-" 
          apply - 
          apply(unfold Is_D_strand_def,unfold  node_sign_def 
            strand_def index_def, auto) 
          done

        from M0  and a8 and M2 and a1
        have M000:"(fst m',1):fst b&((fst m', 1) ,m'):(snd b)^*"
          apply-  apply(simp) 
          apply((erule conjE)+) 
          apply(subgoal_tac "1<= snd m'") 
          apply(fold edges_def nodes_def index_def strand_def)
          apply( drule_tac i="1" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
          apply simp apply simp  
          apply(erule conjE, 
            drule_tac x="((strand  m', 1), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
          apply blast+ apply simp+
          done

        from this and M0 have M3:"((fst m',1),m'): (snd b )^+" 
          apply -  
          apply(erule conjE,drule rtranclD) 
          apply(erule disjE)
          apply(fold strand_def index_def)
          apply( simp add:node_equal) 
          apply(unfold strand_def index_def)
          apply auto
          done
          

        from M01 and M2 have M4:"Key k \<sqsubset>  h" by auto

        from this have M5:"Key k \<sqsubset> Crypt k'  h" 
          by(subgoal_tac "h \<sqsubset>  Crypt k' h", 
            blast dest: subterm_trans, unfold parts_relation_def,
            auto)  

        from this and M2 and M3 and M01 and M000 
        have ?thesis 
          apply- apply(simp only:mem_Collect_eq)
          apply(erule conjE)+
          apply(fold edges_def nodes_def)
          apply(drule_tac x="(fst m', 1)" in spec)
          apply  auto done   
      }
      moreover
      {assume a12:"Is_Cat_strand (?s)"
        from M3  and a8 and  a12 have M0:"snd m'=2" 
          apply -
          apply( unfold Is_Cat_strand_def, unfold Domain_def,
            unfold node_sign_def strand_def index_def,auto)
          apply(case_tac "y=0") apply simp 
          apply (case_tac "y=1" )
          apply simp
          apply auto   
          done
        from a12 and this
        have M1:"EX g h . node_term (fst m',0)=g &
          node_term (fst m', 1)=h &
          node_term m'={|g,h|}" (is "EX g h. ?P g h")
          by (unfold node_term_def, unfold Is_Cat_strand_def
            strand_def index_def , auto)
        then obtain g and h where M01:" ?P g h" 
          by auto
        from a12  have M00:"node_sign (fst m', 0)=-& node_sign (fst m', 1)=-" 
          apply - 
          apply(unfold Is_Cat_strand_def,
            unfold  node_sign_def 
            strand_def index_def , auto)
          done

        from M0 and M01 and M2  
        have M4:"Key k \<sqsubset> {|g,h|}"
          by auto

        from this have "Key k \<sqsubset> g | Key k \<sqsubset> h " 
          apply - apply(unfold parts_relation_def)  
          apply(simp only:parts_insert_MPair parts_Un) 
          apply(simp only:parts_insert2)
          apply auto
          done
        
        moreover
        {assume 
          a20: "Key k \<sqsubset> g" 

          from M0  and a8 and M2 and a1  have M000:"(fst m',0):fst b&((fst m', 0) ,m'):(snd b)^*"
            apply-  apply(simp) 
            apply((erule conjE)+) 
            apply(subgoal_tac "1<= snd m'") 
            apply(fold edges_def nodes_def index_def strand_def)
            apply( drule_tac i="0" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
            apply simp apply simp  
            apply(erule conjE, 
              drule_tac x="((strand  m', 0), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
            apply blast+ apply simp+
            done
            

          from this and M0 have M3:"((fst m',0),m'): (snd b )^+" 
            apply -  
            apply(erule conjE,drule rtranclD) 
            apply(erule disjE)
            apply(fold strand_def index_def)
            apply( simp add:node_equal) 
            apply(unfold strand_def index_def)
            apply auto
            done
            
          from this and M2 and M3 and M01 and a20 and M000  
          have ?thesis
            apply- apply(simp only:mem_Collect_eq)
            apply(erule conjE)+
            apply(fold edges_def nodes_def)
            apply(drule_tac x="(fst m', 0)" in spec)
            apply  auto done   
          
        }   
        moreover
        {assume 
          a20: "Key k \<sqsubset> h" 

          from M0  and a8 and M2 and a1  have M000:"(fst m',1):fst b&((fst m', 1) ,m'):(snd b)^*"
            apply-  apply(simp) 
            apply((erule conjE)+) 
            apply(subgoal_tac "1<= snd m'") 
            apply(fold edges_def nodes_def index_def strand_def)
            apply( drule_tac i="1" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
            apply simp apply simp  
            apply(erule conjE, 
              drule_tac x="((strand  m', 1), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
            apply blast+ apply simp+ 
            done

          from this and M0 
          have M3:"((fst m',1),m'): (snd b )^+" 
            apply -  
            apply(erule conjE,drule rtranclD) 
            apply(erule disjE)
            apply(fold strand_def index_def)
            apply( simp add:node_equal) 
            apply(unfold strand_def index_def)
            apply auto
            done

          from this and M2 and M3 and M01 and a20 and M000 
          have ?thesis 
            apply- apply(simp only:mem_Collect_eq)
            apply(erule conjE)+
            apply(fold edges_def nodes_def)
            apply(drule_tac x="(fst m', 1)" in spec)
            apply  auto done  
          
        }   
        ultimately have ?thesis by blast  
      }


      moreover
      {assume a12:"Is_Sep_strand (?s)"
        from M3 and a8 and  a12 have M0:"snd m'=2|snd m'=1" 
          apply - apply( unfold Is_Sep_strand_def,
            unfold Domain_def, unfold node_sign_def 
            strand_def index_def,auto)
          apply(case_tac "y=0")
          apply simp  
          apply (case_tac "y=1" ) 
          apply simp 
          apply auto   
          done  
        from a12  
        have M00:"node_sign (fst m', 0)=-&
          node_sign (fst m', 1)=+&
          node_sign (fst m', 2)=+" 
          apply - 
          apply(unfold Is_Sep_strand_def,unfold  node_sign_def 
            strand_def index_def, auto)
          done
        from a12 and this 
        have M1:"EX g h . node_term (fst m',0)={|g,h|} &
          node_term (fst m', 1)=g 
          &  node_term (fst m',2)=h" (is "EX g h. ?P g h")
          by (unfold node_term_def, unfold Is_Sep_strand_def
            strand_def index_def, auto)
        then obtain g and h where M01:" ?P g h"  by auto 
        from M0  have M0:"snd m'=2|snd m'=1" by simp
        moreover
        {assume a20:"snd m'=2"

          from M0  and a8 and M2 and a1  have M000:"(fst m',0):fst b&((fst m', 0) ,m'):(snd b)^*"
            apply-  apply(simp) 
            apply((erule conjE)+) 
            apply(subgoal_tac "0<= snd m'") 
            apply(fold edges_def nodes_def index_def strand_def)
            apply( drule_tac i="0" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
            apply simp apply simp  
            apply(erule conjE, 
              drule_tac x="((strand  m', 0), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
            apply blast+ 
            done

          from this and a20  have M3:"((fst m',0),m'): (snd b )^+" 
            apply -  
            apply(erule conjE,drule rtranclD) 
            apply(erule disjE)
            apply(fold strand_def index_def)
            apply( simp add:node_equal) 
            apply(unfold strand_def index_def)
            apply auto
            done
           
          from a20  and M01 have M02: "node_term m'= h" 
            apply - apply ((erule conjE)+,simp) 
            apply(drule sym) 
            apply simp 
            done

          from  M2 and  M01 and this  
          have N0: "Key k \<sqsubset> {| g,h|} "
            apply - 
            apply(simp, subgoal_tac "h \<sqsubset>{| g,h|}")
            apply( blast dest: subterm_trans) 
            apply( unfold parts_relation_def, auto)
            done 

          from this and M2 and M3 and M01 and a20 and M000 and N0  
          have ?thesis
            apply- apply(simp only:mem_Collect_eq)
            apply(erule conjE)+
            apply(fold edges_def nodes_def)
            apply(drule_tac x="(fst m', 0)" in spec)
            apply  auto done  
         
        }
        moreover
        {assume a20:"snd m'=1"

          from M0  and a8 and M2 and a1  
          have M000:"(fst m',0):fst b&((fst m', 0) ,m'):(snd b)^*"
            apply-  apply(simp) 
            apply((erule conjE)+) 
            apply(subgoal_tac "0<= snd m'") 
            apply(fold edges_def nodes_def index_def strand_def)
            apply( drule_tac i="0" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
            apply simp apply simp  
            apply(erule conjE, 
              drule_tac x="((strand  m', 0), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
            apply blast+ 
            done
            

          from this and a20  
          have M3:"((fst m',0),m'): (snd b )^+" 
            apply -  
            apply(erule conjE,drule rtranclD) 
            apply(erule disjE)
            apply(fold strand_def index_def)
            apply( simp add:node_equal) 
            apply(unfold strand_def index_def)
            apply auto
            done

          from a20  and M01 have M02: "node_term m'= g" 
            apply - 
            apply ((erule conjE)+,simp)
            apply(drule sym)
            apply simp done

          from  M2 and  M01 and this  have N0: "Key k \<sqsubset> {| g,h|} "
            apply - 
            apply(simp, subgoal_tac "g \<sqsubset>{| g,h|}")
            apply( blast dest: subterm_trans) 
            apply( unfold parts_relation_def, auto)
            done 

          from this and M2 and M3 and M01 and a20 and M000 and N0 
          have ?thesis 
            apply- apply(simp only:mem_Collect_eq)
            apply(erule conjE)+
            apply(fold edges_def nodes_def)
            apply(drule_tac x="(fst m', 0)" in spec)
            apply  auto done  
          
        }    
        ultimately have ?thesis by blast  
      }

      moreover
      {assume a9:"Is_Tee_strand ?s"
        from a9 and a8 and M3  have M0: "snd m' =2 | snd m'=1" 
          apply - apply( unfold Is_Tee_strand_def,
            unfold Domain_def, 
            unfold node_sign_def 
            strand_def index_def)
          apply(case_tac "snd m'=0") 
          apply( simp)
          apply auto done
        from a9 and this
        have "EX g . node_term (fst m',0)=g & node_term (fst m', 1)=g &  node_term m'=g" (is "EX g. ?P g")
          by (unfold node_term_def, 
            unfold Is_Tee_strand_def 
            strand_def index_def, 
            auto)
        then obtain g where M01:" ?P g" ..
        from a9 have M00:"node_sign (fst m', 0)=-" 
          apply - 
          apply(unfold Is_Tee_strand_def
            ,unfold  node_sign_def
            strand_def index_def, auto)
          done
        have  ?thesis 
        proof - 
          from M0 have M0: "snd m' =2 | snd m'=1" by simp
          moreover 
          {assume M1:" snd m'=2" 
             from M0  and a8 and M2 and a1  
             have M000:"(fst m',0):fst b&((fst m', 0) ,m'):(snd b)^*"
               apply-  apply(simp) 
               apply((erule conjE)+) 
               apply(subgoal_tac "0<= snd m'") 
               apply(fold edges_def nodes_def index_def strand_def)
               apply( drule_tac i="0" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
               apply simp apply simp  
               apply(erule conjE, 
                 drule_tac x="((strand  m', 0), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
               apply blast+ 
               done
            

             from this and M1  
             have M5:"((fst m',0),m'): (snd b )^+" 
               apply -  
               apply(erule conjE,drule rtranclD) 
               apply(erule disjE)
               apply(fold strand_def index_def)
               apply( simp add:node_equal) 
               apply(unfold strand_def index_def)
               apply auto
               done

          

             from this and M2 and M3 and M01 and M1 and M000  
             have ?thesis
               apply- apply(simp only:mem_Collect_eq)
               apply(erule conjE)+
               apply(fold edges_def nodes_def)
               apply(drule_tac x="(fst m', 0)" in spec)
               apply  auto 
               done 
          }
          moreover
          {assume M1:" snd m'=1"
            from M0  and a8 and M2 and a1  
            have M000:"(fst m',0):fst b&((fst m', 0) ,m'):(snd b)^*"
              apply-  apply(simp) 
              apply((erule conjE)+) 
              apply(subgoal_tac "0<= snd m'") 
              apply(fold edges_def nodes_def index_def strand_def)
              apply( drule_tac i="0" and ?n2.0="m'" and b="b" in bundle_casual2_closed3)
              apply simp apply simp  
              apply(erule conjE, 
                drule_tac x="((strand  m', 0), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1)
              apply blast+ 
              done
            

             from this and M1  
             have M5:"((fst m',0),m'): (snd b )^+" 
               apply -  
               apply(erule conjE,drule rtranclD) 
               apply(erule disjE)
               apply(fold strand_def index_def)
               apply( simp add:node_equal) 
               apply(unfold strand_def index_def)
               apply auto
               done

          

             from this and M2 and M3 and M01 and M1 and M000  
             have ?thesis
               apply- apply(simp only:mem_Collect_eq)
               apply(erule conjE)+
               apply(fold edges_def nodes_def)
               apply(drule_tac x="(fst m', 0)" in spec)
               apply  auto 
               done 
            
          }
          ultimately show  ?thesis by blast
        qed}
      moreover
      {assume a9:"Is_T_strand ?s"
        from M3 and a9 and a8  have "snd m'=0"
          apply -
          apply (unfold Is_T_strand_def,
            unfold node_sign_def index_def strand_def,
            unfold Domain_def )
          apply auto 
          done

        from M3 and a9 and a8  
        have "EX t:T. node_term  m'=t" apply -
          apply (unfold Is_T_strand_def,
            unfold node_term_def index_def strand_def,
            unfold Domain_def )
          apply auto 
          done

        
        from this obtain t 
          where  "t:T &node_term  m'=t"
          by auto

        from  this and M2 
        have ?thesis 
          apply - 
          apply(unfold T_is_only_text)
          apply (simp) 
          apply(erule conjE)+  
          apply(drule subterm_atom)
          apply simp   
          apply(drule_tac x="k" in spec)
          apply simp 
          done
      } 
      moreover
      {assume a10 :"Is_K_strand ?s"
        from M3  and a10 and a8  
        have "snd m'=0" apply -
          apply (unfold Is_K_strand_def, unfold node_sign_def 
            index_def strand_def, unfold Domain_def )
          apply auto done

        from M3 and a10 and a8  have "EX k:KP. node_term  m'=Key k" apply -
          apply (unfold Is_K_strand_def, unfold node_term_def, 
            unfold Domain_def index_def strand_def ) 
          apply auto 
          done

        from this obtain k' 
          where M5:" k':KP &  node_term  m'=Key k'"
          by auto
        
        
        
        from  this and M2   have ?thesis 
          apply- 
          apply(simp only:mem_Collect_eq)
          apply(erule conjE)+
          apply(drule subterm_atom) 
          apply(unfold Atoms_def)
          apply simp 
          apply simp
          done
      } 
      moreover
      { assume a11:"Is_Flush_strand ?s"
        from M3 and a11 and a8 have ?thesis
          apply-
          apply(unfold Is_Flush_strand_def,
            unfold Domain_def index_def strand_def ,
            unfold node_sign_def, auto) 
          done
      }


      
      ultimately have  ?thesis by blast
    }

    moreover
    { assume a3:" ~Is_penetrator_strand ?s"
      from a3 
      have M0:" ((EX A B Na Nb. is_initiator (fst m') A B Na Nb) | 
        (EX A B Na Nb. is_responder (fst m') A B Na Nb))"
        apply -
        apply(cut_tac M2 a1)
        apply(subgoal_tac "(strand m'):Σ")
          thm bundle_nodes_strand_in_space
          prefer 2
        apply(blast intro: bundle_nodes_strand_in_space)
        apply(unfold NSL_1 strand_def)
        apply simp 
        done
      
      moreover
      { assume a4:"(EX A B Na Nb. is_initiator (fst m') A B Na Nb)"

        from a4  obtain A  and B and  Na and  Nb 
          where a5:" is_initiator (fst m') A B Na Nb" 
          by auto

        from a5 and M2 and   M3 and a8 
        have "snd m'=0 | snd m'=2"
          apply - 
          apply(unfold is_initiator_def) 
          apply(simp only:mem_Collect_eq) 
          apply(unfold node_sign_def,
            unfold node_term_def,
            unfold Domain_def) 
          apply(case_tac "snd m'=1")
          apply(unfold parts_relation_def)
          apply auto 
          done
        
        from a5 and M2 and M3 and this
        have ?thesis  
          apply - 
          apply(unfold is_initiator_def)
          apply(simp only:mem_Collect_eq) 
          apply( unfold node_term_def) 
          apply(unfold parts_relation_def)
          apply auto 
          done
      }
      moreover
      {  assume a4:"(EX A B Na Nb. is_responder (fst m') A B Na Nb)"
        
        from a4  obtain A  and B and  Na and  Nb 
          where a5:" is_responder (fst m') A B Na Nb" by auto

        from a5 and M2 and   M3 and a8
        have "snd m'=1"
          apply - apply(unfold is_responder_def)
          apply(simp only:mem_Collect_eq) 
          apply(unfold node_sign_def, 
            unfold node_term_def, unfold Domain_def) 
          apply(case_tac "snd m'=2")
          apply(unfold parts_relation_def) 
          apply simp 
          apply(case_tac "snd m'=0") 
          apply auto
          done
        
        from a5 and M2 and M3 and this  
        have ?thesis  
          apply - 
          apply(unfold is_responder_def) 
          apply(simp only:mem_Collect_eq) 
          apply( unfold node_term_def) 
          apply(unfold parts_relation_def)
          apply auto 
          done  
      }
      
      ultimately have ?thesis by blast
    }
    ultimately show ?thesis by blast
  qed
qed

lemma Needham_public_encrypt:
  assumes a1: "M={Crypt (pubK B) {|Nonce Na, Agent A|}} Un {t. ~Nonce Na\<sqsubset>  t} "  
  
  
  shows  c:"~Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|}: synth M"
proof(rule ccontr)
  assume M0: "¬ Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> ∉ synth M"
  show "False"
  proof -
    from M0 have M1:" Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> : synth M"
      by auto
    from this and a1 
    have  M2:"Key (pubK A): synth M & \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>: synth M" 
      apply - 
      apply(unfold parts_relation_def, auto)
      done
    from this and a1 
    have "Nonce Na : synth M & Nonce Nb : synth M " 
      by (unfold parts_relation_def,auto)
    from this 
    have "Nonce Na : synth M & Nonce Nb : synth M " by simp
    from this and a1 
    show ?thesis  
      by (unfold parts_relation_def,auto)
  qed
qed



lemma Needham_public0:  assumes a1: "M={Crypt (pubK B) {|Nonce Na, Agent A|}}
  Un {t. ~Nonce Na\<sqsubset>  t}
  Un {t. (EX r N. (r,1): fst b  & t=node_term (r,1)& is_responder r A B Na N&
  node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})}  "
  and a2:" t: synth M" 
  shows  c:"Nonce Na \<sqsubset> t-->~t={|Nonce Na0, Nonce Nb0, Agent B0|}"
proof
  assume a3:"Nonce Na \<sqsubset> t"
  show c0:"t ≠ \<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>"
  proof(rule ccontr)
    assume a9:"¬ t ≠ {|Nonce Na0, Nonce Nb0, Agent B0|}"  
    show False
    proof -
      from a9 have a10:" t = {|Nonce Na0, Nonce Nb0, Agent B0|}" 
        by auto
      
      from this and a2 have a11:"{|Nonce Na0, Nonce Nb0, Agent B0|}:synth M"
        apply - 
        apply simp
        done
      
      from this    show ?thesis thm MPair_synth
      proof(erule_tac X=" Nonce Na0" and Y="{|Nonce Nb0, Agent B0|}"  
          and P=False in MPair_synth)
        assume a11: "\<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace> ∈ M"
        show  False 
        proof - 
          from a1 and a11 and a3 and a10 
          show ?thesis 
            by auto  
        qed
      next
        assume a12:"Nonce Na0 ∈ synth M" and 
          a13:" \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ synth M"
        show False
        proof -
          from a12 have a14:"Nonce Na0: M" by (simp )
          from a13 have a141:"Nonce Nb0:  M |\<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M" 
          proof(erule_tac  X=" Nonce Nb0" and Y=" Agent B0"  
              and P="Nonce Nb0:  M|\<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M" in MPair_synth) 
            assume a15:" \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M"
            show "Nonce Nb0 ∈ M ∨ \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M" 
            proof - from a15 show ?thesis by auto qed
          next
            assume a15:" Nonce Nb0 ∈ synth M"
            show "Nonce Nb0 ∈ M ∨ \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M"
            proof -
              from a15 have "Nonce Nb0 ∈ M "   by (simp)
              from this show ?thesis by auto
            qed
          qed
          
          from a3 and a10 have M1: "Na0=Na | Nb0=Na" 
            by(unfold parts_relation_def, auto)
          moreover
          {assume a11:"Na0=Na"
            from a11 and a14 and a1 have ?thesis
              apply -
              apply (unfold parts_relation_def,auto ) 
              done
          }
          moreover
          {assume a11:"Nb0=Na"
            from a11 and a141 and a1 have ?thesis 
              apply - 
              apply (unfold parts_relation_def,auto ) 
              done
          }
          ultimately show ?thesis by blast
        qed
      qed
    qed
  qed
qed




lemma Needham_public_encrypt1: 
  assumes a1: "M={Crypt (pubK B) {|Nonce Na, Agent A|}}
  Un {t. ~Nonce Na\<sqsubset>  t}
  Un {t. (EX r N. (r,1): fst b &  t=node_term (r,1)&is_responder r A B Na N&
  node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})}  "
  and a2:" t: synth M" 
  
  
  shows  c:"t=Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}& Nonce Na \<sqsubset> t -->
  (EX r N. (r,1): fst b  & t=node_term (r,1)&is_responder r A B Na N&
  node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})"
proof
  assume a3:" t=Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}&Nonce Na \<sqsubset> t "
  show c0:"(EX r N. (r,1): fst b &  t=node_term (r,1)&is_responder r A B Na N&
    node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})"
  proof -
    from a3 have M0:"Nonce Na \<sqsubset>  {|Nonce Na0, Nonce Nb0, Agent B0|}" 
      apply -  
      apply(erule conjE) 
      apply( simp,drule body_subterm, auto)
      done
    from this  have M00:"Na=Na0 |Na=Nb0" 
      apply - 
      apply(unfold parts_relation_def) 
      apply auto 
      done
    from a2 and a3 
    have M1:"Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}:synth M"
      by simp
    from a1  and this 
    have  M2:"Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}:M
      |Key (pubK A0): synth M & \<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>: synth M" 
      apply -  
      apply( erule Crypt_synth)
      apply( auto)
      done
    moreover
    {assume M3:"Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}:M"
      from M00 and M3 and a1 and a3 have ?thesis 
        apply - 
        apply (simp) 
        apply blast 
        done 
    }
    moreover
    {assume M3:"Key (pubK A0): synth M & \<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>: synth M"
      from  M3 and a1 and a3 and M0  have ?thesis
        apply -  thm Needham_public0
        apply( drule_tac t="\<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>"
          and ?Na0.0="Na0" and ?Nb0.0="Nb0" and
          ?B0.0="B0" in Needham_public0)
        apply blast+ done 
    }
    ultimately show ?thesis by blast
  qed
qed

lemma Needham_public_encrypt2:  assumes  a2:" t: synth ({Crypt (pubK B) {|Nonce Na, Agent A|}}
  Un {t. ~Nonce Na\<sqsubset>  t}
  Un {t. (EX r N. (r,1): fst b  & t=node_term (r,1)&is_responder r A B Na N&
  node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})})" 
  
  
  shows  c:"t=Crypt (pubK B0) {|Nonce Na0, Agent A0|}& Nonce Na \<sqsubset> t -->
  t=Crypt (pubK B) {|Nonce Na, Agent A|}"
proof
  assume a3:"t=Crypt (pubK B0) {|Nonce Na0, Agent A0|}&Nonce Na \<sqsubset> t "
  show c0:"t=Crypt (pubK B) {|Nonce Na, Agent A|}"
  proof -
    from a3 have M0:"Nonce Na \<sqsubset>   {|Nonce Na0, Agent A0|}" 
      apply -  
      apply(erule conjE) 
      apply( simp,drule body_subterm, auto) 
      done
    from this  have M00:"Na=Na0" 
      apply - 
      apply(unfold parts_relation_def) 
      apply auto
      done

    from a2 and a3
    have M1:"Crypt (pubK B0) {|Nonce Na0, Agent A0|}:
      synth ({Crypt (pubK B) {|Nonce Na, Agent A|}}
      Un {t. ~Nonce Na\<sqsubset>  t}
      Un {t. (EX r N. (r,1): fst b  & t=node_term (r,1)&is_responder r A B Na N&
      node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})})" by simp
    let ?M="{Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪ {t. ¬ Nonce Na \<sqsubset>  t} ∪
      {t. ∃r N. (r, 1::'a) ∈ fst b ∧  t = node_term (r, 1)
      ∧is_responder r A B Na N&
      node_term (r, 1) = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>}"

    from M1 have M01:"Crypt (pubK B0) {|Nonce Na0, Agent A0|}:synth ?M"
      apply - apply auto done

    from M1 and a3 have  M2:"B0 = B ∧ Na0 = Na ∧ A0 = A
      |Key (pubK B0): synth ?M & \<lbrace>Nonce Na0, Agent A0\<rbrace>: synth ?M" 
      apply - thm Crypt_synth
      apply(erule_tac ?K="(pubK B0)" and ?X="{|Nonce Na0, Agent A0|}" and
        ?H="?M" 
        and  ?P="B0 = B ∧ Na0 = Na ∧ A0 = A
        |Key (pubK B0): synth ?M & \<lbrace>Nonce Na0, Agent A0\<rbrace>: synth ?M" in Crypt_synth) 
      apply( auto) done
    moreover
    {assume M3:"B0 = B ∧ Na0 = Na ∧ A0 = A"
      from M3 and a3 have ?thesis  apply - apply (simp)  done 
    }
    moreover
    {assume M3:"Key (pubK B0): synth ?M & \<lbrace>Nonce Na0,  Agent A0\<rbrace>: synth ?M"
      from  M3   have "Nonce Na0 :synth ?M | \<lbrace>Nonce Na0,  Agent A0\<rbrace>:?M"
        apply - apply(erule conjE,erule MPair_synth) apply auto done

      moreover
      {assume M4:"Nonce Na0 :synth ?M"
        from this and M00 have ?thesis 
          apply -
          apply simp 
          apply(unfold parts_relation_def) 
          apply auto 
          done
      }
      
      moreover
      {assume M4:"\<lbrace>Nonce Na0,  Agent A0\<rbrace>:?M"
        from this and M00 have ?thesis 
          apply - 
          apply simp 
          apply(unfold parts_relation_def)  
          apply auto 
          done
      }
      ultimately have ?thesis by blast
    }
    ultimately show ?thesis by blast
    
    
    
  qed
qed

lemma is_initiator_unique_by_Na:
  "[|is_initiator (fst m) A0 B0 Na0 Nb0;  snd m = 1; 
  node_term m = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>|]
  ==> A0=A & Na0=Na & Nb0=N &B0=B"
  apply(unfold is_initiator_def,unfold node_term_def, auto)
  done

lemma initiator_auth: 
  assumes a1:"b: bundles" and 
  a2:" is_initiator s A B Na Nb" and 
  a3:" (s,2): nodes b" and 
  a4:"~A:bad" and a5:"~B:bad" and a6:"~Na=Nb"  and
  a7:"unique_originate (Nonce Na) (s,0)"
  shows "EX r. is_responder r A B Na Nb & (r,1):nodes b"
proof -
  from a2 have M1:"node_term (s,0)= Crypt (pubK B) {|Nonce Na, Agent A|} &
    node_term  (s,1)=Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|}&
    node_term (s,2)= Crypt (pubK B) (Nonce Nb)&
    node_sign (s,0)=+&
    node_sign  (s,1)=-&
    node_sign (s,2)=+ " 
    apply- 
    apply(unfold is_initiator_def, 
      unfold node_term_def, 
      unfold node_sign_def
    )
    apply auto
    done

  from a5 
  have M2: " ~ invKey ( pubK B): KP " 
    apply -
    apply(unfold KP_def)
    apply( auto)
    done
  
  
  from this  have M3: "ALL  n: (fst b). ~node_term n = Key (invKey ( pubK B)) "  thm  ccontr
  proof(rule_tac P="ALL n: fst b. ~node_term n = Key (invKey ( pubK B)) " in ccontr)
    assume M4:"¬ (∀n∈fst b. node_term n ≠ Key (priEK B))"
    from M4 have M5:"(EX n : fst b. node_term n = Key (priEK B))" by auto
    
    from this obtain n where M6:"n : nodes b &  node_term n = Key (priEK B)" 
      by(unfold nodes_def, auto)

    from a1 and M2 and M6 
    have M7:" ~node_term n = Key (priEK B)"
      by(blast dest: no_node_is_key) 

    from this and M6 
    show False 
      by simp
  qed  

  from a4 
  have M21: " ~ invKey ( pubK A): KP "
    by(unfold KP_def, auto)
  
  
  from this  
  have M31: "ALL  n: (fst b). ~node_term n = Key (invKey ( pubK A)) "  thm  ccontr
  proof(rule_tac P="ALL n: fst b. ~node_term n = Key (invKey ( pubK A)) " in ccontr)
    assume M4:"¬ (∀n∈fst b. node_term n ≠ Key (priEK A))"
    from M4 have M5:"(EX n : fst b. node_term n = Key (priEK A))" by auto
    
    from this obtain n where M6:"n : nodes b &  node_term n = Key (priEK A)"
      by (unfold nodes_def,auto)

    from a1 and M21 and M6 have M7:" ~node_term n = Key (priEK A)" 
      by(blast dest: no_node_is_key) 

    from this and M6 
    show False 
      by simp
  qed  
  

  let ?M= "{Crypt (pubK B) {|Nonce Na, Agent A|}} Un {t. ~Nonce Na\<sqsubset>  t} Un
    {t. (EX r N. (r,1): fst b  & t=node_term (r,1)&is_responder r A B Na N&
    node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})}
    "
  let  ?a="Nonce Na"

  have  N0:"Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|}: synth  ?M"
  proof(rule ccontr)
    assume c0:"Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>
      ∉ synth ({Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪ {t. ¬ Nonce Na \<sqsubset>  t} ∪
      {t. ∃r N. (r, 1) ∈ fst b  ∧
      t = node_term (r, 1) ∧is_responder r A B Na N&
      node_term (r, 1) = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>})"
    

    from M3 and M31  
    have b10:"ALL t:  ?M. ?a\<sqsubset> t --> (ALL k h. t = Crypt k h -->
      (ALL n:fst b. node_term n = Key (invKey k) --> Is_regular_strand (fst n)))"
      apply -
      apply auto  
      done
    

    have b4:"ALL t: ?M. ?a\<sqsubset> t--> (ALL g h. ~ t ={|g, h|})" 
      apply - apply auto done

    have b11:"?a:Atoms & (ALL k. ~?a=Key k)" 
      apply - 
      apply (unfold Atoms_def,auto) 
      done

    have b3:"ALL t.~?a\<sqsubset> t -->t:?M" by blast

    from b3 b4 b10 have b10:"suite ?M ?a b"
      apply - apply(unfold suite_def)
      apply( unfold component_def regularK_def 
        nodes_def strand_def)
      apply(rule conjI) 
      prefer 2
      apply blast
      apply(rule ballI)
      apply(rule impI)+
      apply(rule conjI)
      apply(rotate_tac 1)
      apply(drule_tac x="t" in bspec)
      apply blast apply blast
      apply(rotate_tac 2)
      apply(rule allI)+
      apply(drule_tac x="t" in bspec)
      apply blast
      apply(rule impI)
      apply(rule ballI) 
      
      apply blast
      done
    
    let ?n="(s,0)" let ?n'="(s,1)" 
    
    from a1 and a3 have a02:"?n':nodes b" 
      apply - 
      apply (drule_tac i="1" in  bundle_casual2_closed3) 
      apply (unfold strand_def index_def,auto)
      done
    from  a7
    have  a03:" unique_originate ?a ?n" 
      by auto
    from M1 have  a7a:" node_sign ?n'=-"
      by auto

    
    from M1 and c0 
    have  b12: " ?a\<sqsubset>node_term ?n' & ~  node_term ?n': (synth ?M)" 
      apply - apply auto  done

    from a1 and a02 have a8:"?n':Domain"
      by(blast dest: bundle_node_in_Domain)

    from a1 and a2 and a6  
    have a10:" ALL  n0. strand n0=strand ?n& node_sign n0=+& 
      ?a \<sqsubset> node_term n0 & n0: nodes b--> (node_term n0: (synth ?M))"  
      apply- apply(unfold is_initiator_def, unfold node_term_def,
        unfold node_sign_def strand_def nodes_def )
      apply auto
      apply(subgoal_tac "(s, ba):Domain") 
      prefer 2
      apply(fold nodes_def )
      apply(blast dest: bundle_node_in_Domain)
      apply(unfold Domain_def strand_def
        index_def,auto)
      apply(case_tac "ba=0") 
      apply auto  
      apply(case_tac "ba=1" )
      apply auto 
      apply(case_tac "ba=2") 
      apply auto
      apply(subgoal_tac "¬ Nonce Na \<sqsubset> Crypt (pubK B) (Nonce Nb)")
      apply simp
      apply(unfold parts_relation_def) 
      apply(simp add:parts_insert_Crypt)
      done

    from M1 have "node_term ?n:synth ?M"
     by blast

     
     
    with a1 and b10 and b11 and b4 and b3 and b12 and a02 and a03 and a8 and a10 and a7a

    have M10:"EX  m' m.((?n,m):(edges b)^* & (m',?n'):(edges b)^*&(m, m'):casual2^+ 
              & node_sign m'=+ 
              &( strand m'~=strand ?n &  node_sign m=- )
              & ?a \<sqsubset> node_term m & ?a \<sqsubset> node_term m'
              &  ~  node_term m': (synth ?M) &  node_term m: (synth ?M)
              &Is_regular_strand (strand m')
              &(ALL y.(y,m'):(edges b)^+ -->~ ?a\<sqsubset> node_term y| (node_term y):(synth ?M))) " 
      (is "EX m' m. ?P m' m")
      apply - apply(rule_tac n="?n" in  general_auth_test1)
      apply assumption+
      apply(simp) apply blast+  
      done


    from this obtain p and m' and m where M11:"?P m' m" 
      apply - 
      apply(erule exE)+ 
      apply auto 
      done

   

   
    from a1 a02  M11 have  M161:"m':nodes b" thm bundle_is_rtrans_closed23
      by (blast dest:bundle_is_rtrans_closed23)
    
    from this and a1  have M17a:" m':Domain" 
      by (blast dest:bundle_node_in_Domain)

    from a1  and M11 M161 have M16:" m: nodes b" 
      apply -
      apply(blast dest:bundle_casual2_trancl)
      done
    
    from this and a1  have M17:" m:Domain"
      by (blast dest:bundle_node_in_Domain)



    from M11 have M170:"fst m'= fst m & snd m < snd m'"  
      by (fold index_def strand_def,blast dest: casul2_trancl_sign)

    have M172:"?M= {Crypt (pubK B) {|Nonce Na, Agent A|}} Un {t. ~Nonce Na\<sqsubset>  t} Un
      {t. (EX r N. (r,1): fst b  & t=node_term (r,1)&is_responder r A B Na N&
      node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})}"
      by simp

    from M11  M161 have M18:"((EX A B Na Nb. is_initiator (fst m') A B Na Nb)
      | (EX A B Na Nb. is_responder (fst m') A B Na Nb))"
      apply -
      apply(cut_tac a1)
      apply(subgoal_tac "(strand m'):Σ")
      thm bundle_nodes_strand_in_space
      prefer 2
      apply(fold strand_def)
      apply(blast intro: bundle_nodes_strand_in_space)
      apply(unfold NSL_1 Is_regular_strand_def  strand_def)
      apply simp 
      done

    moreover
    { assume c20:"EX A B Na Nb. is_initiator (fst m') A B Na Nb"
      from this obtain A0 B0 Na0 Nb0 where c21:"is_initiator (fst m') A0 B0 Na0 Nb0" by auto
      from M11 and c21 and M17  have M13:"snd m=1"
        apply - apply(unfold is_initiator_def)      
        apply(case_tac "snd m=1") apply simp
        apply(subgoal_tac "fst m'= fst m")
        prefer 2 
        apply(fold strand_def)
        apply(blast dest: casul2_trancl_sign)  
        apply simp
        apply (unfold node_sign_def strand_def index_def) 
        apply(case_tac "snd m=0") apply simp 
        apply(case_tac "snd m=2") apply simp
        apply(unfold Domain_def) apply auto
        done
      from c21 and this and M170 
      have M19:"node_term m= Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}"
        apply -  
        apply(unfold is_initiator_def, unfold node_term_def)
        apply auto 
        done 
      thm Needham_public_encrypt1
      from M172 and M11 and this 
      have M20:"(∃r N. (r, 1) ∈ fst b ∧  node_term m = node_term (r, 1)
        ∧is_responder r A B Na N& node_term (r, 1) = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>)" 
        (is "EX r N. ?P1 r N") 
        apply - 
        apply ((erule conjE)+) thm  Needham_public_encrypt1
        apply(drule_tac M="({Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪ {t. ¬ Nonce Na \<sqsubset>  t} ∪
          {t. ∃r N. (r, 1) ∈ fst b ∧
          t = node_term (r, 1) ∧is_responder r A B Na N
          & node_term (r, 1) = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>})" 
          and b="b" and t="node_term m" and Na="Na" and ?A0.0="A0" 
          and ?Na0.0="Na0" and ?Nb0.0="Nb0" in  Needham_public_encrypt1)
        apply blast+
        done
      from this obtain r N 
        where M21:"?P1 r N" 
        by auto
      from this have M22:"node_term m=Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>"
        by simp 
      from c21 and M170 and M13 and this
      have  M22:"A0=A & Na0=Na & Nb0=N &B0=B" 
        by(simp, blast dest:is_initiator_unique_by_Na)
      from this and c21  have M23:"is_initiator (fst m') A B Na N" 
        by simp
      from this and a2 and a7 have M24:"N=Nb"
        by (blast dest:unique_orig_init)
      from this and M23 
      have M25:" is_initiator (fst m') A B Na Nb" 
        by simp
      from M170 and M13 and M17a and M25
      have "snd m'=2 & node_term m'=Crypt (pubK B) (Nonce Nb)"
        apply - 
        apply(unfold is_initiator_def,unfold Domain_def,
          unfold node_term_def strand_def index_def) 
        apply(case_tac "snd m'=0")  apply simp
        apply(case_tac "snd m'= 1") apply simp
        apply(case_tac "snd m'= 2") apply( auto) 
        done
      from M11 and this have "Na=Nb"
        apply -  
        apply(erule conjE)+ 
        thm body_subterm 
        apply(simp, drule_tac ?t''="Nonce Na" in body_subterm) 
        apply simp 
        apply (unfold parts_relation_def,simp) 
        done
      from this and a6 
      have ?thesis
        by auto
    }
    
    moreover
    { assume c20:"EX A B Na Nb. is_responder (fst m') A B Na Nb"
      from this obtain A0 B0 Na0 Nb0 
        where c21:"is_responder (fst m') A0 B0 Na0 Nb0" by auto
      from M11 and c21 and M17a  have M13:"snd m'=1"
        apply - apply(unfold is_responder_def)      
        apply(case_tac "snd m'=1") apply simp
        apply (unfold node_sign_def) 
        apply(case_tac "snd m'=0") apply simp 
        apply(case_tac "snd m'=2") apply simp
        apply(unfold Domain_def) apply auto
        done 

      from M170 and this and M17 have M18:"snd m=0"
        apply - apply((erule conjE)+)
        apply(case_tac "snd m'=0",auto) 
        done
      
      from c21 and this and M170 
      have M19:"node_term m= Crypt (pubK B0) {|Nonce Na0,  Agent A0|}"
        apply -  
        apply(unfold is_responder_def, unfold node_term_def)
        apply auto  
        done

      from c21 and M13 and M170 
      have M20:"node_term m'= Crypt (pubK A0) {|Nonce Na0, Nonce Nb0 ,Agent B0|}"
        apply -  
        apply(unfold is_responder_def, unfold node_term_def) 
        apply auto  
        done
      
      from M19 and M11 
      have  M21:"node_term m=Crypt (pubK B) {|Nonce Na, Agent A|}"
        apply - 
        apply ((erule conjE)+) 
        apply(drule Needham_public_encrypt2) 
        apply auto
        done
      
      from M19 and this 
      have M22:"B0=B& Na0=Na& A0=A" 
        by simp

      from M13 have M22a:"m'=(fst m',1)"
        apply  (cut_tac surjective_pairing[where p="m'"]) 
        apply simp 
        done

      from M161 and M13 and this  
      have M220:"(fst m',1): fst b" 
        apply - 
        apply(fold nodes_def, auto)
        done

     
      from this and M220 and M22a and M20 and c21 and M22
      have M23:"node_term m': {t. ∃r N. (r, 1) ∈ fst b 
        ∧ t = node_term (r, 1)&is_responder r A B Na N 
        ∧ node_term (r, 1) = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>}"
        apply - apply (simp only:mem_Collect_eq)
        apply(rule_tac x="fst m'" in exI, rule_tac x="Nb0" in exI)
        apply auto done
      
      from this have "node_term m': ?M" by blast

      from this have "node_term m': synth ?M" by blast

      from this and M11 have ?thesis by blast
    }
    ultimately have ?thesis  by blast
    
    from this and c0 show False by auto
  qed
  have M172:"?M= {Crypt (pubK B) {|Nonce Na, Agent A|}} Un {t. ~Nonce Na\<sqsubset>  t} Un
    {t. (EX r N. (r,1): fst b  & t=node_term (r,1)&is_responder r A B Na N&
    node_term (r,1)=Crypt (pubK A) {|Nonce Na, Nonce N, Agent B|})}" 
    by simp
  from this and N0  
  have "(∃r N . (r, 1) ∈ fst b ∧
    Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>=node_term (r,1)&
    is_responder r A B Na N& node_term (r, 1) = 
    Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>)" 
    (is "EX r N . ?P1 r N ") 
    apply - 
    apply(drule_tac M="({Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪ {t. ¬ Nonce Na \<sqsubset>  t} ∪
      {t. ∃r N. (r, 1) ∈ fst b ∧
      t = node_term (r, 1) ∧is_responder r A B Na N
      & node_term (r, 1) = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>})" 
      and b="b" and t="Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>" 
      and Na="Na" and ?A0.0="A"  and ?B0.0="B" 
      and  ?Na0.0="Na" and ?Nb0.0="Nb" in 
      Needham_public_encrypt1)
    apply simp
    apply(subgoal_tac "Nonce Na \<sqsubset> Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>")
    apply(drule mp)
    apply blast
    apply simp 
    apply (unfold parts_relation_def, auto)
    done

  from this obtain r N where N1:"?P1 r N" by auto 

  from this have N2:"N=Nb" 
    apply- 
    apply ( unfold is_responder_def,unfold node_term_def, auto) done

  from this and N1 have N2:"is_responder r A B Na Nb" by simp

  from this and N1 show ?thesis 
    by(fold nodes_def, blast)
qed


lemma Needham_public0:  assumes a1: "M={Crypt (pubK B) (Nonce Nb), 
            Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|} } Un {t. ~Nonce Nb\<sqsubset>  t} "
  and a2:" t: synth M" 
  shows  c:"Nonce Nb \<sqsubset> t-->~t={|Nonce Na0, Nonce Nb0, Agent B0|}"
proof
  assume a3:"Nonce Nb \<sqsubset> t"
  show c0:"t ≠ \<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>"
  proof(rule ccontr)
    assume a9:"¬ t ≠ {|Nonce Na0, Nonce Nb0, Agent B0|}"  
    show False
    proof -
      from a9 have a10:" t = {|Nonce Na0, Nonce Nb0, Agent B0|}" by auto
      
      from this and a2 have a11:"{|Nonce Na0, Nonce Nb0, Agent B0|}:synth M" apply - apply simp done
      
      from this    show ?thesis thm MPair_synth
      proof(erule_tac X=" Nonce Na0" and Y="{|Nonce Nb0, Agent B0|}"  and P=False in MPair_synth)
        assume a11: "\<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace> ∈ M"
        show  False 
        proof - from a1 and a11 and a3 and a10 show ?thesis by auto  qed
      next
        assume a12:"Nonce Na0 ∈ synth M" and a13:" \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ synth M"
        show False
        proof -
          from a12 have a14:"Nonce Na0: M" by (simp )
          from a13 have a141:"Nonce Nb0:  M |\<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M" 
          proof(erule_tac  X=" Nonce Nb0" and Y=" Agent B0"  and P="Nonce Nb0:  M|\<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M" in MPair_synth) 
            assume a15:" \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M"
            show "Nonce Nb0 ∈ M ∨ \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M" 
            proof - from a15 show ?thesis by auto qed
          next
            assume a15:" Nonce Nb0 ∈ synth M"
            show "Nonce Nb0 ∈ M ∨ \<lbrace>Nonce Nb0, Agent B0\<rbrace> ∈ M"
            proof -
              from a15 have "Nonce Nb0 ∈ M "   by (simp)
              from this show ?thesis by auto
            qed
          qed
          
          from a3 and a10 have M1: "Na0=Nb | Nb0=Nb"  by(unfold parts_relation_def, auto)
          moreover
          {assume a11:"Na0=Nb"
            from a11 and a14 and a1 have ?thesis apply - apply (unfold parts_relation_def,auto ) done
          }
          moreover
          {assume a11:"Nb0=Nb"
            from a11 and a141 and a1 have ?thesis apply - apply (unfold parts_relation_def,auto ) done
          }
          ultimately show ?thesis by blast
        qed
      qed
    qed
  qed
qed

lemma Needham_public_encrypt1:  assumes a1: "M={Crypt (pubK B) (Nonce Nb), 
            Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|} } Un {t. ~Nonce Nb\<sqsubset>  t} "
  and a2:" t: synth M" 
  
  
  shows  c:"t=Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}& Nonce Nb \<sqsubset> t -->
        A0=A&B0=B&Na0=Na&Nb0=Nb"
 proof
   assume a3:"t = Crypt (pubK A0) \<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace> ∧ Nonce Nb \<sqsubset>  t"
   show "A0=A&B0=B&Na0=Na&Nb0=Nb"
   proof -
     from a3 have M0:"Nonce Nb \<sqsubset> {|Nonce Na0, Nonce Nb0, Agent B0|}"
       apply -  apply(erule conjE) apply( simp,drule body_subterm, auto) done
     from this  have M00:"Nb=Na0 |Nb=Nb0" apply - apply(unfold parts_relation_def) apply auto done
     from a2 and a3 have M1:"Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}:synth M" by simp
     from a1  and this have  M2:"Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}:M
      |Key (pubK A0): synth M & \<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>: synth M" 
       apply -  apply( erule Crypt_synth) apply( auto) done
     moreover
     {assume M3:"Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}:M"
       from M00 and M3 and a1 and a3 have ?thesis  apply - apply (simp) apply blast done 
     }
     moreover
     {assume M3:"Key (pubK A0): synth M & \<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>: synth M"
       from  M3 and a1 and a3 and M0  have ?thesis
         apply -  thm Needham_public0
         apply( drule_tac t="\<lbrace>Nonce Na0, Nonce Nb0, Agent B0\<rbrace>" and
           ?Na0.0="Na0" and ?Nb0.0="Nb0" and ?B0.0="B0" in Needham_public0)
        apply blast+ done 
    }
    ultimately show ?thesis by blast
  qed
qed

lemma Needham_public_encrypt2:  assumes  a1: "M={Crypt (pubK B) (Nonce Nb), 
            Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|} } Un {t. ~Nonce Nb\<sqsubset>  t} "
  and a2:" t: synth M"  
  
  shows  c:"  Nonce Nb \<sqsubset> t -->
  ( Crypt (pubK B) (Nonce Nb)\<sqsubset> t | Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|} \<sqsubset> t )" 
    (is "(?P t)")
  using a2
proof(induct)
  fix X
  assume b1:"X:M"
  show "(?P X)"
  proof 
    assume b2:"Nonce Nb \<sqsubset>  X" 
    show "Crypt (pubK B) (Nonce Nb) \<sqsubset>  X ∨ Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> \<sqsubset>  X"
    proof -
      from a1 b1 b2 show ?thesis
        apply - apply(unfold parts_relation_def, auto)
        done
    qed
  qed
next
  fix agt
  show "?P (Agent agt)" by (unfold parts_relation_def,auto)
next
  fix n
  show "?P (Number n)" by (unfold parts_relation_def,auto)
next
  fix X
  assume b1:"X:synth M" and b2:"Nonce Nb \<sqsubset>  X -->
    Crypt (pubK B) (Nonce Nb) \<sqsubset>  X ∨ Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> \<sqsubset>  X"
  show "?P (Hash X)"
    proof
      assume b3:" Nonce Nb \<sqsubset>  Hash X"
      show "Crypt (pubK B) (Nonce Nb) \<sqsubset>  Hash X ∨ Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> \<sqsubset>  Hash X"
        proof -
          from b3 have False
            by(unfold parts_relation_def, auto)
          
          then show "Crypt (pubK B) (Nonce Nb) \<sqsubset>  Hash X ∨
            Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> \<sqsubset>  Hash X" by auto
        qed
      qed
next
  fix X Y
  assume b1:"X ∈ synth M" and 
  b2:"?P X" and b3:"?P Y"
  show "?P {|X,Y|}"
    proof
      assume b1:"Nonce Nb \<sqsubset>  \<lbrace>X, Y\<rbrace>"
      show " Crypt (pubK B) (Nonce Nb) \<sqsubset>  \<lbrace>X, Y\<rbrace> ∨ Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> \<sqsubset>  \<lbrace>X, Y\<rbrace>"
        proof -
          from b1 have "Nonce Nb \<sqsubset> X | Nonce Nb \<sqsubset>  Y"
            apply - apply(unfold parts_relation_def,simp add:parts_insert_MPair parts_insert2) 
            done
         with b2 b3 show ?thesis
              apply - apply(unfold parts_relation_def, simp add:parts_insert_MPair parts_insert2 )
              apply auto
              done
          qed
        qed
next
  fix K X
  assume b1:"X ∈ synth M" and 
  b2:"?P X" and b3:"Key K ∈ M"
  show "?P (Crypt K X)"
    proof
      assume b1:"Nonce Nb \<sqsubset> Crypt K X "
      show " Crypt (pubK B) (Nonce Nb) \<sqsubset>  Crypt K X ∨ 
        Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> \<sqsubset>  Crypt K X"
        proof -
          from b1 have "Nonce Nb \<sqsubset> X "
            apply - apply(unfold parts_relation_def,simp add:parts_insert_Crypt parts_insert2) 
            done
          with b2 b3 show ?thesis
            apply - apply(unfold parts_relation_def, simp add:parts_insert_MPair parts_insert2 )
            apply auto
            done
        qed
      qed
qed



lemma secrecy_of_nb:
assumes a1:"b: bundles" and
  a2:"is_responder r A B Na Nb"  and 
  a3:" (r,2): nodes b" and 
  a4:"~A:bad"  and
  a5:"~B:bad" and
  a6:"~Na=Nb"  and a7:"unique_originate (Nonce Nb) (r,1)"
  shows "ALL n:(nodes b). node_term n ~=Nonce Nb"
proof 
  fix n
  assume M0:"n:nodes b"
  show "node_term n ~=Nonce Nb"
    proof -
       from a2 have M1:"node_term (r,0)= Crypt (pubK B) {|Nonce Na, Agent A|} &
         node_term  (r,1)=Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|}&
         node_term (r,2)= Crypt (pubK B) (Nonce Nb)&
         node_sign (r,0)=-&
         node_sign  (r,1)=+&
         node_sign (r,2)=- " 
         apply- 
         apply(unfold  is_responder_def) 
         apply( unfold node_term_def,
           unfold node_sign_def strand_def index_def)
         apply( auto) 
         done

       from a5 
       have M2: " ~ invKey ( pubK B): KP " 
         apply -
         apply(unfold KP_def)
         apply( auto)
         done
  
  
      from this  
      have M3: "ALL  n: (nodes b). ~node_term n = Key (invKey ( pubK B)) " 
      proof(rule_tac 
          P="ALL n: nodes b. ~node_term n = Key (invKey ( pubK B)) " 
          in ccontr)
        assume M4:"¬ (∀n∈nodes b. node_term n ≠ Key (priEK B))"
        from M4 have M5:"(EX n : nodes b. node_term n = Key (priEK B))" 
          by auto
    
        from this obtain n 
          where M6:"n : nodes b &  node_term n = Key (priEK B)" 
          by(unfold nodes_def,auto)

        from a1 and M2 and M6 
        have M7:" ~node_term n = Key (priEK B)" 
          by(blast dest: no_node_is_key) 

        from this and M6 show False by simp
      qed  

      from a4 
      have M21: " ~ invKey ( pubK A): KP " 
        by(unfold KP_def, auto)
  
  
      from this  have M31: "ALL  n: (fst b). ~node_term n = Key (invKey ( pubK A)) "  thm  ccontr
      proof(rule_tac P="ALL n: fst b. ~node_term n = Key (invKey ( pubK A)) " in ccontr)
        assume M4:"¬ (∀n∈fst b. node_term n ≠ Key (priEK A))"
        from M4 
        have M5:"(EX n : fst b. node_term n = Key (priEK A))" 
          by auto
    
        from this obtain n where 
          M6:"n : fst b &  node_term n = Key (priEK A)" 
          by auto

        from a1 and M21 and M6 
        have M7:" ~node_term n = Key (priEK A)" 
          by(fold nodes_def,blast dest: no_node_is_key) 

        from this and M6 show False by simp
      qed  
  

      let ?M= "{Crypt (pubK B) (Nonce Nb), 
            Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|} }
            Un {t. ~Nonce Nb\<sqsubset>  t}
        "
      let  ?a="Nonce Nb"

      have  N0:"?a \<sqsubset> node_term n --> node_term n: synth  ?M" 
      proof(rule impI,rule ccontr)
        assume c0:"~node_term n: synth ?M" and c1:"?a \<sqsubset> node_term n"
        show False
          proof -

            from M3 and M31 
            have b10:"ALL t:  ?M. ?a\<sqsubset> t --> 
              (ALL k h. t = Crypt k h --> 
              (ALL n:nodes b. node_term n = Key (invKey k) -->Is_regular_strand (strand n) ))"
              apply - 
              apply(fold nodes_def, auto)
              done

            have b4:"ALL t: ?M. ?a\<sqsubset> t--> (ALL g h. ~ t ={|g, h|})" 
              apply - 
              apply auto
              done

            
    
            let ?n="(r,1)" let ?n'="n" 
    
           
            

    
            from  c0 c1 
            have  b12:  " ?a\<sqsubset>node_term n & ~  node_term n: (synth ?M)" 
              apply - apply simp  done

           (* from a1 and a02 have a8:"?n':Domain" by(blast dest: bundle_node_in_Domain)*)

            from a1 a2 
            have a10:" ALL  n0. strand  n0=strand ?n& node_sign n0=+& ?a \<sqsubset> node_term n0 & n0: nodes b
              --> (node_term n0: (synth ?M))"
              apply- apply(unfold is_responder_def) 
              apply( unfold node_sign_def node_term_def nodes_def strand_def )
              apply auto
              apply(subgoal_tac "(r, ba):Domain") 
              prefer 2 
              apply(fold nodes_def strand_def )
              apply(blast dest: bundle_node_in_Domain)        
              apply(unfold Domain_def,auto)
              apply(unfold nodes_def strand_def )
              apply(case_tac "ba=0")
              apply auto  
              apply(case_tac "ba=1" )
              apply auto  
              apply(case_tac "ba=2")
              apply auto 
              done
            have b3:"ALL t.~?a\<sqsubset> t -->t:?M" by blast 
            
            from b3 b4 b10 have b10:"suite ?M ?a b"
              apply - apply(unfold suite_def)
              apply( unfold component_def regularK_def 
                nodes_def strand_def)
              apply(rule conjI) 
              prefer 2
              apply blast
              apply(rule ballI)
              apply(rule impI)+
              apply(rule conjI)
              apply(rotate_tac 1)
              apply(drule_tac x="t" in bspec)
              apply blast apply blast
              apply(rotate_tac 2)
              apply(rule allI)+
              apply(drule_tac x="t" in bspec)
              apply blast
              apply(rule impI)
              apply(rule ballI) 
      
              apply blast
              done

            from a1 and b10  and b4 and b3 and b12  and a7  and a10 and M1

          
            have M10:"EX  m' m.((?n,m):(edges b)^* & (m',?n'):(edges b)^*&(m, m'):casual2^+ 
              & node_sign m'=+ 
              &( strand m'~=strand ?n &  node_sign m=- )
              & ?a \<sqsubset> node_term m & ?a \<sqsubset> node_term m'
              &  ~  node_term m': (synth ?M) &  node_term m: (synth ?M)
              &Is_regular_strand (strand m')
              &(ALL y.(y,m'):(edges b)^+ -->~ ?a\<sqsubset> node_term y| (node_term y):(synth ?M))) " 
              (is "EX m' m. ?P m' m")


              apply - apply(rule_tac n="?n" in   general_auth_test1)
              apply assumption+
              apply(simp) 
              apply blast+ 
              done


            from this obtain m' and m where M11:"?P m' m " 
              apply - 
              apply(erule exE)+ 
              apply auto 
              done

            from a1 M0  M11 have  M161:"m':nodes b" 
              thm bundle_is_rtrans_closed23
              by (blast dest:bundle_is_rtrans_closed23)
    
            from this and a1  
            have M17a:" m':Domain" 
              by (blast dest:bundle_node_in_Domain)

            from a1  and M11 M161 
            have M16:" m: nodes b" 
              apply -
              apply(blast dest:bundle_casual2_trancl)
              done
    
            from this and a1  
            have M17:" m:Domain"
              by (blast dest:bundle_node_in_Domain)



            from M11 
            have M170:"fst m'= fst m & snd m < snd m'"  
              by (fold index_def strand_def,
                blast dest: casul2_trancl_sign)

            have M172: "?M=?M" by simp
           
            from M11  M161 have M18:"((EX A B Na Nb. is_initiator (fst m') A B Na Nb)
              | (EX A B Na Nb. is_responder (fst m') A B Na Nb))"
              apply -
              apply(cut_tac a1)
              apply(subgoal_tac "(strand m'):Σ")
              thm bundle_nodes_strand_in_space
              prefer 2
              apply(fold strand_def)
              apply(blast intro: bundle_nodes_strand_in_space)
              apply(unfold NSL_1  Is_regular_strand_def strand_def)
              apply simp 
              done


            moreover
            { assume c20:"EX A B Na Nb. is_initiator (fst m') A B Na Nb"
              from this obtain A0 B0 Na0 Nb0 where c21:"is_initiator (fst m') A0 B0 Na0 Nb0" by auto
              from M11 and c21 and M17  have M13:"snd m=1"
                apply - apply(unfold is_initiator_def)      
                apply(case_tac "snd m=1") apply simp
                apply(subgoal_tac "fst m'= fst m")
                prefer 2 
                apply(fold strand_def)
                apply(blast dest: casul2_trancl_sign)
                apply simp 
                apply (unfold strand_def index_def node_sign_def) 
                apply(case_tac "snd m=0") apply simp 
                apply(case_tac "snd m=2") apply simp
                apply(unfold Domain_def) apply auto
                done
              from c21 and this and M170
              have M19:"node_term m= 
                Crypt (pubK A0) {|Nonce Na0, Nonce Nb0, Agent B0|}"
                apply -  
                apply(unfold is_initiator_def, unfold node_term_def)
                apply auto  
                done 
              thm Needham_public_encrypt1
              from M172 M11 this 
              have M25:"A0=A&B0=B&Na0=Na&Nb0=Nb"
                apply - apply ((erule conjE)+)
                apply(drule  Needham_public_encrypt1) 
                apply blast+
                done

              
              from this and c21  
              have M23:"is_initiator (fst m') A B Na Nb" 
                by simp
              
              from M170 and M13 and M17a c21 M25  
              have "snd m'=2 & node_term m'=Crypt (pubK B) (Nonce Nb)"
                apply - 
                apply(unfold is_initiator_def,unfold Domain_def,
                  unfold node_term_def strand_def index_def) 
                apply(case_tac "snd m'=0")  apply simp
                apply(case_tac "snd m'= 1") apply simp 
                apply(case_tac "snd m'= 2") apply( auto) 
                done
              then  have "node_term m':synth ?M" 
                by auto
              with M11 have ?thesis 
                by auto
            }

    
            moreover
            { assume c20:"EX A B Na Nb. is_responder (fst m') A B Na Nb"
              from this obtain A0 B0 Na0 Nb0 where c21:"is_responder (fst m') A0 B0 Na0 Nb0" by auto
              from M11 and c21 and M17a  have M13:"snd m'=1"
                apply - apply(unfold is_responder_def strand_def index_def)      
                apply(case_tac "snd m'=1") apply simp
                apply (unfold node_sign_def) 
                apply(case_tac "snd m'=0") apply simp 
                apply(case_tac "snd m'=2") apply simp
                apply(unfold Domain_def) apply auto
                done 

              from M170 and this and M17 have M18:"snd m=0"
                apply - apply((erule conjE)+)
                apply(case_tac "snd m'=0",auto) 
                done
      
              from c21 and this and M170 
              have M19:"node_term m=
                Crypt (pubK B0) {|Nonce Na0,  Agent A0|}"
                apply -  
                apply(unfold is_responder_def, unfold node_term_def)
                apply auto 
                done
              
              from c21 and M13 and M170 
              have M20:"node_term m'= Crypt (pubK A0) {|Nonce Na0, Nonce Nb0 ,Agent B0|}"
                apply -  
                apply(unfold is_responder_def, unfold node_term_def) 
                apply auto 
                done

              have "?M=?M" by simp
              from this  M11  
              have  M21:"( Crypt (pubK B) (Nonce Nb)\<sqsubset> node_term m 
                | Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|} \<sqsubset> node_term m )"
                apply -
                apply ((erule conjE)+)
                apply(drule Needham_public_encrypt2)
                apply auto 
                done

              moreover
              {assume M22:"Crypt (pubK B) (Nonce Nb)\<sqsubset> node_term m"
                from M22 M19  have False
                  by (unfold parts_relation_def,auto)
                then have ?thesis
                  by simp
              }
              moreover
              {assume M22:"Crypt (pubK A) {|Nonce Na, Nonce Nb, Agent B|}\<sqsubset> node_term m"
                from M22 M19  have False
                  by (unfold parts_relation_def,auto)
                then have ?thesis
                  by simp
              }
              ultimately  have ?thesis by blast
            }
            ultimately  show ?thesis by blast

          qed
        qed
        have "Nonce Nb \<sqsubset> node_term n | ~Nonce Nb \<sqsubset> node_term n" 
          by simp
        moreover
        {assume N1:"Nonce Nb \<sqsubset> node_term n"
          from N0 N1  have "node_term n :synth ?M" by simp
          then have "node_term n ~=Nonce Nb" 
            by(unfold parts_relation_def, auto)
        }
        moreover
        {assume N1:"~Nonce Nb \<sqsubset> node_term n"
          from N0  have "node_term n ~=Nonce Nb" 
            by(unfold parts_relation_def, auto)
        }
        ultimately show ?thesis by blast
    qed
qed
end




lemma unique_orig_init:

  [|is_initiator s A B Na Nb; is_initiator s' A' B' Na Nb';
   unique_originate (Nonce Na) (s, 0)|]
  ==> s = s'A = A'B = B'Nb = Nb'

lemma c:

  [|b ∈ bundles; n ∈ nodes b|] ==> ∀k. k ∉ KP --> node_term n ≠ Key k

lemma c:

  M = {Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪
      {t. ¬ Nonce Na \<sqsubset>  t} ==>
  Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> ∉ synth M

lemma c:

  [|M = {Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪
        {t. ¬ Nonce Na \<sqsubset>  t} ∪
        {t. ∃r N. (r, 1::'a) ∈ fst bt = node_term (r, 1) ∧
                  is_responder r A B Na N ∧
                  node_term (r, 1) =
                  Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>};
   t ∈ synth M|]
  ==> Nonce Na \<sqsubset>  t -->
      t ≠ \<lbrace>Nonce Na0.0, Nonce Nb0.0, Agent B0.0\<rbrace>

lemma c:

  [|M = {Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪
        {t. ¬ Nonce Na \<sqsubset>  t} ∪
        {t. ∃r N. (r, 1::'a) ∈ fst bt = node_term (r, 1) ∧
                  is_responder r A B Na N ∧
                  node_term (r, 1) =
                  Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>};
   t ∈ synth M|]
  ==> t = Crypt (pubK A0.0)
           \<lbrace>Nonce Na0.0, Nonce Nb0.0, Agent B0.0\<rbrace> ∧
      Nonce Na \<sqsubset>  t -->
      (∃r N. (r, 1::'a) ∈ fst bt = node_term (r, 1) ∧
             is_responder r A B Na N ∧
             node_term (r, 1) =
             Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>)

lemma c:

  t ∈ synth
       ({Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>} ∪
        {t. ¬ Nonce Na \<sqsubset>  t} ∪
        {t. ∃r N. (r, 1::'a) ∈ fst bt = node_term (r, 1) ∧
                  is_responder r A B Na N ∧
                  node_term (r, 1) =
                  Crypt (pubK A)
                   \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>}) ==>
  t = Crypt (pubK B0.0) \<lbrace>Nonce Na0.0, Agent A0.0\<rbrace> ∧
  Nonce Na \<sqsubset>  t -->
  t = Crypt (pubK B) \<lbrace>Nonce Na, Agent A\<rbrace>

lemma is_initiator_unique_by_Na:

  [|is_initiator (fst m) A0.0 B0.0 Na0.0 Nb0.0; snd m = 1;
   node_term m = Crypt (pubK A) \<lbrace>Nonce Na, Nonce N, Agent B\<rbrace>|]
  ==> A0.0 = ANa0.0 = NaNb0.0 = NB0.0 = B

lemma initiator_auth:

  [|b ∈ bundles; is_initiator s A B Na Nb; (s, 2) ∈ nodes b; A ∉ bad; B ∉ bad;
   NaNb; unique_originate (Nonce Na) (s, 0)|]
  ==> ∃r. is_responder r A B Na Nb ∧ (r, 1) ∈ nodes b

lemma c:

  [|M = {Crypt (pubK B) (Nonce Nb),
         Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>} ∪
        {t. ¬ Nonce Nb \<sqsubset>  t};
   t ∈ synth M|]
  ==> Nonce Nb \<sqsubset>  t -->
      t ≠ \<lbrace>Nonce Na0.0, Nonce Nb0.0, Agent B0.0\<rbrace>

lemma c:

  [|M = {Crypt (pubK B) (Nonce Nb),
         Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>} ∪
        {t. ¬ Nonce Nb \<sqsubset>  t};
   t ∈ synth M|]
  ==> t = Crypt (pubK A0.0)
           \<lbrace>Nonce Na0.0, Nonce Nb0.0, Agent B0.0\<rbrace> ∧
      Nonce Nb \<sqsubset>  t -->
      A0.0 = AB0.0 = BNa0.0 = NaNb0.0 = Nb

lemma c:

  [|M = {Crypt (pubK B) (Nonce Nb),
         Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace>} ∪
        {t. ¬ Nonce Nb \<sqsubset>  t};
   t ∈ synth M|]
  ==> Nonce Nb \<sqsubset>  t -->
      Crypt (pubK B) (Nonce Nb) \<sqsubset>  t ∨
      Crypt (pubK A) \<lbrace>Nonce Na, Nonce Nb, Agent B\<rbrace> \<sqsubset>  t

lemma secrecy_of_nb:

  [|b ∈ bundles; is_responder r A B Na Nb; (r, 2) ∈ nodes b; A ∉ bad; B ∉ bad;
   NaNb; unique_originate (Nonce Nb) (r, 1)|]
  ==> ∀n∈nodes b. node_term n ≠ Nonce Nb