Up to index of Isabelle/HOL/strand
theory otway_rees05theory otway_rees05= strand05+ Public: constdefs is_initiator::"sigma => agent=> agent =>nat =>msg=>key =>bool" "is_initiator s A B Na M K == (SP s)=[(+, {|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|}|}), (-, {|M,Crypt (shrK A) {|Nonce Na, Key K|} |})] & M:T" constdefs is_responder::"sigma => agent=> agent =>nat =>msg=>key=>msg=>msg =>bool" "is_responder s A B Nb M K H H' == (SP s)=[(-, {|M, Agent A, Agent B, H|}), (+, {|M, Agent A, Agent B, H, Crypt (shrK B) {|Nonce Nb, M, Agent A, Agent B|} |}), (-, {|M, H', Crypt (shrK B) {|Nonce Nb,Key K|}|} ), (+, {|M,H'|} )]& M:T&(~M=Nonce Nb)& (~ Nonce Nb \<sqsubset> H) " constdefs is_server::"sigma => agent=> agent =>nat =>nat=>msg=>key =>bool" "is_server s A B Na Nb M K== (ALL ag. ~K=shrK ag)& (SP s)=[ (-, {|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|}, Crypt (shrK B) {|Nonce Nb, M, Agent A, Agent B|} |}), (+, {|M, Crypt (shrK A) {|Nonce Na, Key K|}, Crypt (shrK B) {|Nonce Nb,Key K|}|})]& M:T " defs OR_1: "Σ == {s. (Is_penetrator_strand s) | (EX A B Na M k. is_initiator s A B Na M k) | (EX A B Nb M K H H'. is_responder s A B Nb M K H H')| (EX A B Na Nb M K. is_server s A B Na Nb M K) }" lemma no_node_is_key: assumes a1:"b:bundles" and a2:"n: nodes b" shows c:"ALL ag. ( ~ag: bad) --> ~node_term n=Key (shrK ag)" proof (rule_tac P="ALL ag. ( ~ag: bad) --> ~node_term n=Key (shrK ag)" in ccontr) assume a3:"~(ALL ag. ( ~ag: bad) --> ~node_term n=Key (shrK ag)) " show "False" proof - from a3 have M0:"EX ag. ( ~ag:bad )&node_term n = Key (shrK ag)" by auto from M0 obtain ag where b1: "( ~ag :bad) & node_term n = Key (shrK ag)" by auto let ?k=" (shrK ag)" from b1 have b2: "( ~?k :KP) " apply - apply(unfold KP_def, blast ) done let ?P="{x. Key ?k \<sqsubset> node_term x & x:nodes b }" 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''" by (unfold casual1_def,auto) 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 - 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)" by (unfold penetrator_trace,auto) 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 and b1 have ?thesis apply- apply(simp only:mem_Collect_eq) apply(erule conjE)+ apply(drule subterm_atom) apply(unfold Atoms_def KP_def) apply simp apply auto 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 M k. is_initiator (fst m') A B Na M k) | (EX A B Nb M K H H'. is_responder (fst m') A B Nb M K H H') |(EX A B Na Nb M K. is_server (fst m') A B Na Nb M K))" 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 OR_1 strand_def) apply simp done moreover { assume a4:"(EX A B Na M k. is_initiator (fst m') A B Na M k)" from a4 obtain A and B and Na and M and k where a5:" (is_initiator (fst m') A B Na M k)" by auto from a5 and M2 and M3 and a8 have "snd m'=0 " 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 this have "node_term m'={|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|}|}& M:T" by(unfold is_initiator_def, unfold node_term_def, auto) from this and M2 have ?thesis apply - apply( unfold node_term_def) apply(unfold T_is_only_text) apply(unfold parts_relation_def) apply(unfold Atoms_def, auto) done } moreover { assume a4:"(EX A B N M K H H'. is_responder (fst m') A B N M K H H')" from a4 obtain A and B and N and M and K and H and H' where a5:" is_responder(fst m') A B N M K H H'" by auto from a5 have a90: "node_term (fst m',0)={|M, Agent A, Agent B, H|}& node_term (fst m',1)={|M, Agent A, Agent B, H, Crypt (shrK B) {|Nonce N, M, Agent A, Agent B|} |}& node_term (fst m',2)= {|M, H', Crypt (shrK B) {|Nonce N,Key K|}|} & node_term (fst m',3)= {|M,H'|}& M:T" apply - apply(unfold is_responder_def, unfold node_term_def, auto) done from a5 and M2 and M3 and a8 have a9:"snd m'=1 | snd m'=3" 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(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply simp apply(case_tac "snd m'=3") apply simp apply(case_tac "snd m'=2") apply auto done moreover { assume a91:"snd m'=1" from a91 and M2 and a90 have a10:" Key ?k \<sqsubset> H" apply - apply(unfold parts_relation_def, auto) apply(subgoal_tac "m'=(fst m',1)") apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(erule disjE) apply simp apply(fold parts_relation_def) apply simp apply(drule subterm_atom) apply auto apply(simp add:node_equal ) apply(unfold strand_def index_def,auto) done from a10 and a90 have a12:"Key ?k \<sqsubset> node_term (fst m',0)" by (unfold parts_relation_def, simp add:parts_insert2) from a1 and a91 and M2 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^*" apply - apply (simp only:mem_Collect_eq, (erule conjE)+) apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) 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 a91 have N3:"(fst m',0): fst b & ((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 a12 and M2 and M3 and N3 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 a91:"snd m'=3" from a91 and M2 and a90 have a10:" Key ?k \<sqsubset> H'" apply - apply(unfold parts_relation_def, auto) apply(subgoal_tac "m'=(fst m',3)") apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(erule disjE) apply simp apply(fold parts_relation_def) apply(drule subterm_atom) apply auto apply(simp add:node_equal) apply (unfold index_def strand_def, auto) done from a10 and a90 have a12:"Key ?k \<sqsubset> node_term (fst m',2)" by (unfold parts_relation_def, simp add:parts_insert2) from a1 and a91 and M2 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^*" apply - apply (simp only:mem_Collect_eq, (erule conjE)+) apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="2" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply simp apply(erule conjE, drule_tac x="((strand m', 2), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1) apply blast+ done from this and a91 have N3:"(fst m',2): fst b & ((fst m',2),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 a12 and M2 and N3 have ?thesis apply- apply(simp only:mem_Collect_eq) apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 2)" in spec) apply auto done } ultimately have ?thesis by blast } moreover { assume a5:"EX A B Na Nb M K.(is_server (fst m') A B Na Nb M K)" from a5 obtain A and B and Na and Nb and M and K where a5:" (is_server (fst m') A B Na Nb M K)" by auto from a5 and M2 and M3 and a8 have "snd m'=1 " apply - apply(unfold is_server_def) apply(simp only:mem_Collect_eq) apply(unfold node_sign_def, unfold node_term_def index_def strand_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from a5 and this have "node_term m'={|M, Crypt (shrK A) {|Nonce Na, Key K|}, Crypt (shrK B) {|Nonce Nb,Key K|}|}" apply - apply(unfold is_server_def, unfold node_term_def, auto) done from this and M2 and a5 have ?thesis apply - apply( unfold node_term_def, unfold is_server_def) apply(unfold parts_relation_def) apply(auto) apply(unfold T_is_only_text) apply(unfold Atoms_def, auto) done } ultimately have ?thesis by blast } ultimately show False by blast qed qed lemma server_auth: assumes a1:"b: bundles" and a2:" is_server s A B Na Nb M K" and a3:" (s,1): nodes b" and a4:"~A:bad" and a6:"~A=B" shows " (EX s Ks. is_initiator s A B Na M Ks & (s,0):nodes b ) " proof - from a2 have a7: "node_term (s,0)= {|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|}, Crypt (shrK B) {|Nonce Nb, M, Agent A, Agent B|} |}& node_term (s,1)={|M, Crypt (shrK A) {|Nonce Na, Key K|}, Crypt (shrK B) {|Nonce Nb,Key K|}|} " apply - apply(unfold is_server_def, unfold node_term_def, auto) done from a1 and a3 have a8:"(s,0): fst b & ((s,0), (s,1)): (edges b)^*" apply - apply (drule_tac i="0" and ?n2.0="(s,1)" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((s, 0), (s,1)) ∈ (edges b )^*" ) apply(drule_tac R="edges b" in rtranclD) apply (erule disjE) apply( simp add:node_equal ) apply(unfold strand_def index_def nodes_def) apply auto apply(blast dest:rtrancl_mono1) done let ?n="(s,0)" let ?k="shrK A" let ?h="{|Nonce Na, M, Agent A, Agent B|}" from a1 and a4 have a9:"ALL n:(nodes b). node_term n = Key ?k --> Is_regular_strand (fst n)" proof(rule_tac A="nodes b" in ballI) fix x assume a10:"b:bundles" and a11:"x:nodes b" and a11a:"~A:bad" show " node_term x = Key (shrK A) --> Is_regular_strand (fst x) " proof assume a12:" node_term x = Key (shrK A)" show "Is_regular_strand (fst x)" proof - from a11a and a10 and a11 have a13:" ~node_term x=Key (shrK A)" by(blast dest:no_node_is_key) from a12 and a13 show ?thesis by auto qed qed qed from a7 have a10:"Crypt ?k ?h \<sqsubset> node_term ?n" apply - apply(unfold parts_relation_def, simp add:parts_insert2) done thm unsolicited_test from a1 and a8 and a9 and a10 have a11:"EX n'. (n',?n):(edges b)^*&(n': (nodes b) & Crypt ?k ?h \<sqsubset> node_term n' & Is_regular_strand (strand n') & node_sign n'=+& (ALL y.(y,n'):(edges b)^+ -->~Crypt ?k ?h \<sqsubset> node_term y))" (is "EX n'.?P n'") apply - apply(fold nodes_def strand_def) apply(rule unsolicited_test) apply auto apply(unfold regularK_def) apply assumption done from a11 obtain m' where a12:"?P m'" by blast from this and a1 have M17a:" m':Domain" by (blast dest:bundle_node_in_Domain) from this a12 have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') |(EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" apply - apply(subgoal_tac "(strand m'):Σ") thm bundle_nodes_strand_in_space prefer 2 apply(fold strand_def) apply(cut_tac a1 ) apply(blast intro: bundle_nodes_strand_in_space) apply(unfold OR_1 Is_regular_strand_def strand_def) apply auto done have a14:"EX K0. is_initiator (fst m') A B Na M K0&(fst m',0):fst b " proof - from a13 have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') |(EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" by simp moreover {assume a15:"(EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0)" from this obtain A0 and B0 and Na0 and M0 and K0 where a16:" is_initiator (fst m') A0 B0 Na0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have b17:"snd m'=0" apply - apply(unfold is_initiator_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def strand_def index_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|}&M0:T" by(unfold is_initiator_def, unfold node_term_def, auto) from a12 and this have "A0=A & B0=B& Na0=Na & M0=M " apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(erule disjE,blast dest: subterm_atom) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this and a16 and b17 and a12 have ?thesis apply - apply(rule_tac x="K0" in exI) apply (simp,subgoal_tac "m'=(fst m',0)") apply(fold nodes_def) apply(erule conjE, simp ) apply(simp add:node_equal) apply(unfold strand_def index_def) apply auto done } moreover {assume a15:"(EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')" from this obtain A0 and B0 and N0 and M0 and K0 and H0 and H0' where a16:" is_responder (fst m') A0 B0 N0 M0 K0 H0 H0'" by blast from a16 have a17: "node_term (fst m',0)={|M0, Agent A0, Agent B0, H0|}& node_term (fst m',1)={|M0, Agent A0, Agent B0, H0, Crypt (shrK B0) {|Nonce N0, M0, Agent A0, Agent B0|} |}& node_term (fst m',2)= {|M0, H0', Crypt (shrK B0) {|Nonce N0,Key K0|}|} & node_term (fst m',3)= {|M0,H0'|}& M0:T" apply - apply(unfold is_responder_def, unfold node_term_def, auto) done from a1 and a12 have a18:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and a16 and a18 have a9:"snd m'=1 | snd m'=3" apply - apply(unfold is_responder_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply simp apply(case_tac "snd m'=3") apply simp apply(case_tac "snd m'=2") apply auto done moreover { assume a91:"snd m'=1" from a91 have "m'=(fst m',1)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0" apply - apply(unfold parts_relation_def) apply (simp add:parts_insert2) apply auto apply(unfold T_is_only_text) apply simp apply(unfold Atoms_def , auto) done from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',0)" by (unfold parts_relation_def, simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^*" apply - apply ( (erule conjE)+) apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) 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 a91 have N3:"(fst m',0): fst b & ((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 b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 0)" in spec) apply auto done from this have ?thesis by blast } moreover { assume a91:"snd m'=3" from a91 have "m'=(fst m',3)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0'" apply - apply(unfold parts_relation_def) apply (simp add:parts_insert2) apply auto apply(unfold T_is_only_text) apply simp apply(unfold Atoms_def , auto) done from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',2)" by (unfold parts_relation_def,simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^* " apply - apply ( (erule conjE)+) apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="2" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply simp apply(erule conjE, drule_tac x="((strand m', 2), m')" and r="(edges b ∩ casual2)" and s="edges b" in rtrancl_mono1) apply blast+ done from this and a91 have N3:"(fst m',2): fst b & ((fst m',2),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 a12 and b12 and N3 and a91 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 2)" in spec) apply auto done then have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a15:" (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0)" from this obtain A0 and B0 and Na0 and Nb0 and M0 and K0 where a16:" is_server (fst m') A0 B0 Na0 Nb0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have "snd m'=1" apply - apply(unfold is_server_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'= {|M0,Crypt (shrK A0) {|Nonce Na0, Key K0|}, Crypt (shrK B0) {|Nonce Nb0,Key K0|}|}&M0:T" by(unfold is_server_def, unfold node_term_def, auto) from a12 and this have "False " apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this and a16 have ?thesis by blast } ultimately show ?thesis by blast qed from this show ?thesis by(fold nodes_def, blast) qed lemma server_auth1: assumes a1:"b: bundles" and a2:" is_server s A B Na Nb M K" and a3:" (s,1): nodes b" and a4:"~A:bad" and a6:"A=B" shows " (EX s Ks. is_initiator s A B Na M Ks & (s,0):nodes b ) | (EX r K H H'. is_responder r A B Na M K H H' & (r,1):nodes b ) " proof - from a2 have a7: "node_term (s,0)= {|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|}, Crypt (shrK B) {|Nonce Nb, M, Agent A, Agent B|} |}& node_term (s,1)= {|M, Crypt (shrK A) {|Nonce Na, Key K|}, Crypt (shrK B) {|Nonce Nb,Key K|}|} " apply - apply(unfold is_server_def, unfold node_term_def, auto) done from a1 and a3 have a8:"(s,0): fst b & ((s,0), (s,1)): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="(s,1)" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this have a8:"(s,0): fst b & ((s,0),(s,1)): (snd b )^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done let ?n="(s,0)" let ?k="shrK A" let ?h="{|Nonce Na, M, Agent A, Agent B|}" from a1 and a4 have a9:"ALL n:(nodes b). node_term n = Key ?k --> Is_regular_strand (strand n)" proof(rule_tac ballI) fix x assume a10:"b:bundles" and a11:"x:nodes b" and a11a:"~A:bad" show " node_term x = Key (shrK A) -->Is_regular_strand (strand x)" proof assume a12:" node_term x = Key (shrK A)" show " Is_regular_strand (strand x)" proof - from a11a and a10 and a11 have a13:" ~node_term x=Key (shrK A)" by(blast dest:no_node_is_key) from a12 and a13 show ?thesis by auto qed qed qed from a7 have a10:"Crypt ?k ?h \<sqsubset> node_term ?n" apply - apply(unfold parts_relation_def, simp add:parts_insert2) done from a1 and a8 and a9 and a10 have a11:"EX n'.(n',?n):(edges b)^*& (n': (nodes b) & Crypt ?k ?h \<sqsubset> node_term n' & Is_regular_strand (strand n') & node_sign n'=+& (ALL y.(y,n'):(edges b)^+ -->~Crypt ?k ?h \<sqsubset> node_term y))" (is "EX n'.?P n'") apply - apply(fold nodes_def strand_def) apply(rule unsolicited_test) apply auto apply(unfold regularK_def) apply assumption done from a11 obtain m' where a12:"?P m'" by blast from this have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') |(EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" apply - apply(subgoal_tac "(strand m'):Σ") thm bundle_nodes_strand_in_space prefer 2 apply(fold strand_def) apply(cut_tac a1 ) apply(blast intro: bundle_nodes_strand_in_space) apply(unfold OR_1 Is_regular_strand_def strand_def) apply auto done show ?thesis proof - from a13 have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') |(EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" by simp moreover {assume a15:"(EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0)" from this obtain A0 and B0 and Na0 and M0 and K0 where a16:" is_initiator (fst m') A0 B0 Na0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have b17:"snd m'=0" apply - apply(unfold is_initiator_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def strand_def index_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|}&M0:T" by(unfold is_initiator_def , unfold node_term_def, auto) from a12 and this have "A0=A & B0=B& Na0=Na & M0=M " apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(erule disjE,blast dest: subterm_atom) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this and a16 and b17 and a12 have ?thesis apply - apply(rule disjI1) apply(unfold nodes_def) apply(rule_tac x="(fst m')" in exI,rule_tac x="K0" in exI) apply(fold nodes_def) apply (simp,subgoal_tac "m'=(fst m',0)") apply(erule conjE, simp ) apply(simp add:node_equal) apply(unfold strand_def index_def) apply(auto) done } moreover {assume a15:"(EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')" from this obtain A0 and B0 and N0 and M0 and K0 and H0 and H0' where a16:" is_responder (fst m') A0 B0 N0 M0 K0 H0 H0'" by blast from a16 have a17: "node_term (fst m',0)={|M0, Agent A0, Agent B0, H0|}& node_term (fst m',1)={|M0, Agent A0, Agent B0, H0, Crypt (shrK B0) {|Nonce N0, M0, Agent A0, Agent B0|} |}& node_term (fst m',2)= {|M0, H0', Crypt (shrK B0) {|Nonce N0,Key K0|}|} & node_term (fst m',3)= {|M0,H0'|}& M0:T" apply - apply(unfold is_responder_def, unfold node_term_def, auto) done from a1 and a12 have a18:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and a16 and a18 have a9:"snd m'=1 | snd m'=3" apply - apply(unfold is_responder_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_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'=3") apply simp apply(case_tac "snd m'=2") apply auto done moreover { assume a91:"snd m'=1" from a91 have "m'=(fst m',1)" apply - apply(simp add:node_equal) apply(unfold strand_def index_def) apply auto done from this and a17 and a12 and a6 have b10:" Crypt ?k ?h \<sqsubset> H0 |(N0=Na&M0=M&B0=A&A0=B)" apply - apply(unfold parts_relation_def) apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(unfold Atoms_def, auto) done moreover {assume b10:" Crypt ?k ?h \<sqsubset> H0" from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',0)" by (unfold parts_relation_def,simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 0)" in spec) apply auto done from this have ?thesis by blast } moreover {assume b10:" (N0=Na&M0=M&B0=A&A0=B)" from a91 have "m'=(fst m',1)" apply(simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b10 and a16 and a12 and a6 and this have ?thesis apply - apply(rule disjI2,rule_tac x="(fst m')"in exI) apply(rule_tac x="K0" in exI, rule_tac x="H0" in exI, rule_tac x="H0'" in exI, auto) done } ultimately have ?thesis by blast } moreover { assume a91:"snd m'=3" from a91 have "m'=(fst m',3)" apply(simp add:node_equal) apply(unfold strand_def index_def) apply auto done from this and a17 and a12 and a6 have b10:" Crypt ?k ?h \<sqsubset> H0'" apply - apply(unfold parts_relation_def, auto) apply(subgoal_tac "m'=(fst m',3)") apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(erule disjE) apply simp apply(fold parts_relation_def) apply( drule subterm_atom) apply(unfold Atoms_def, auto) done from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',2)" by (unfold parts_relation_def,simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="2" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from a12 and b12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 2)" in spec) apply auto done then have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a15:" (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0)" from this obtain A0 and B0 and Na0 and Nb0 and M0 and K0 where a16:" is_server (fst m') A0 B0 Na0 Nb0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have "snd m'=1" apply - apply(unfold is_server_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'= {|M0,Crypt (shrK A0) {|Nonce Na0, Key K0|}, Crypt (shrK B0) {|Nonce Nb0,Key K0|}|}&M0:T" by(unfold is_server_def, unfold node_term_def, auto) from a12 and this have "False " apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this and a16 have ?thesis by blast } ultimately show ?thesis by blast qed qed lemma server_auth_init: assumes a1:"b: bundles" and a2:" is_server s A B Na Nb M K" and a3:" (s,1): nodes b" and a4:"~A:bad" shows " (EX s Ks. is_initiator s A B Na M Ks & (s,0):nodes b ) | (EX r K H H'. is_responder r A B Na M K H H' & (r,1):nodes b ) " proof - from a1 and a2 and a3 and a4 show ?thesis apply - apply(case_tac "A=B") thm server_auth1 apply(drule server_auth1) apply simp+ apply(drule server_auth) apply auto done qed lemma unique_orig_init: assumes a1:"is_initiator s A B Na M K" and a2:"is_initiator s' A' B' Na M' K'" and a3:" unique_originate (Nonce Na) (s,0)" shows "s=s'& A=A'& B=B'&M=M'&K=K'" proof - have "Nonce Na \<sqsubset> \<lbrace>M', Agent A', Agent B', Crypt (shrK A') \<lbrace>Nonce Na, M', Agent A', Agent B'\<rbrace>\<rbrace>" apply - apply(unfold parts_relation_def, auto) apply(simp add:parts_insert2) done from a2 a3 this have "originate (Nonce Na) (s',0)" apply - apply(unfold is_initiator_def originate_def node_sign_def node_term_def parts_relation_def index_def) apply auto done from a1 this a2 a3 show ?thesis apply - apply(unfold is_initiator_def unique_originate_def) apply(erule conjE)+ apply(drule_tac x="(s',0)" in spec) apply simp done qed lemma server_auth_resp0: assumes a1:"b: bundles" and a2:" is_server s A B Na Nb M K" and a3:" (s,1): nodes b" and a5:"~B:bad" and a6:"~A=B" shows " (EX r K H H'. is_responder r A B Nb M K H H' & (r,1):nodes b ) " proof - from a2 have a7: " node_term (s,0)= {|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|}, Crypt (shrK B) {|Nonce Nb, M, Agent A, Agent B|} |}& node_term (s,1)={|M, Crypt (shrK A) {|Nonce Na, Key K|}, Crypt (shrK B) {|Nonce Nb,Key K|}|} " apply - apply(unfold is_server_def,unfold node_term_def, auto) done from a1 and a3 have a8:"(s,0): nodes b & ((s,0), (s,1)): (edges b)^*" apply - apply(drule_tac i="0" and ?n2.0="(s,1)" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this have a8:"(s,0): nodes b & ((s,0),(s,1)): (edges b )^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done let ?n="(s,0)" let ?k="shrK B" let ?h="{|Nonce Nb, M, Agent A, Agent B|}" from a1 and a5 have a9:"ALL n:(nodes b). node_term n = Key ?k -->Is_regular_strand (strand n)" proof(rule_tac A="nodes b" in ballI) fix x assume a10:"b:bundles" and a11:"x:nodes b" and a11a:"~B:bad" show " node_term x = Key (shrK B) -->Is_regular_strand (strand x) " proof assume a12:" node_term x = Key (shrK B)" show "Is_regular_strand (strand x)" proof - from a11a and a10 and a11 have a13:" ~node_term x=Key (shrK B)" by(blast dest:no_node_is_key) from a12 and a13 show ?thesis by auto qed qed qed from a7 have a10:"Crypt ?k ?h \<sqsubset> node_term ?n" apply - apply(unfold parts_relation_def, simp add:parts_insert2) done from a1 and a8 and a9 and a10 have a11:"EX n'.(n',?n):(edges b)^*& (n': (nodes b) & Crypt ?k ?h \<sqsubset> node_term n' & Is_regular_strand (strand n') & node_sign n'=+& (ALL y.(y,n'):(edges b)^+ -->~Crypt ?k ?h \<sqsubset> node_term y))" (is "EX n'.?P n'") apply - apply(rule unsolicited_test) apply auto apply(unfold regularK_def) apply assumption done from a11 obtain m' where a12:"?P m'" by blast from this have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') | (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" apply - apply(subgoal_tac "(strand m'):Σ") thm bundle_nodes_strand_in_space prefer 2 apply(fold strand_def) apply(cut_tac a1 ) apply(blast intro: bundle_nodes_strand_in_space) apply(unfold OR_1 Is_regular_strand_def strand_def) apply auto done have a14:"EX K H H'. is_responder(fst m') A B Nb M K H H' & (fst m',1):fst b " proof - from a13 have a13: " ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') | (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" by simp moreover {assume a15:"(EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0)" from this obtain A0 and B0 and Na0 and M0 and K0 where a16:" is_initiator (fst m') A0 B0 Na0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have b17:"snd m'=0" apply - apply(unfold is_initiator_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'= {|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|}&M0:T" by(unfold is_initiator_def, unfold node_term_def, auto) from a12 and this have "A0=A & B0=B& Na0=Nb & M0=M &A0=B" apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(erule disjE,blast dest: subterm_atom) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this have "A=B" by auto from this and a6 have False by blast then have ?thesis by blast } moreover {assume a15:"(EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')" from this obtain A0 and B0 and N0 and M0 and K0 and H0 and H0' where a16:" is_responder (fst m') A0 B0 N0 M0 K0 H0 H0'" by blast from a16 have a17: "node_term (fst m',0)={|M0, Agent A0, Agent B0, H0|}& node_term (fst m',1)={|M0, Agent A0, Agent B0, H0, Crypt (shrK B0) {|Nonce N0, M0, Agent A0, Agent B0|} |}& node_term (fst m',2)= {|M0, H0', Crypt (shrK B0) {|Nonce N0,Key K0|}|} & node_term (fst m',3)= {|M0,H0'|}& M0:T" apply - apply(unfold is_responder_def, unfold node_term_def, auto) done from a1 and a12 have a18:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and a16 and a18 have a9:"snd m'=1 | snd m'=3" apply - apply(unfold is_responder_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply simp apply(case_tac "snd m'=3") apply simp apply(case_tac "snd m'=2") apply auto done moreover { assume a91:"snd m'=1" from a91 have a92:"m'=(fst m',1)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0 | (B0= B & M0=M &A0=A& N0=Nb) " apply - apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply((erule conjE)+,erule disjE) apply simp+ apply(erule disjE) apply simp apply(fold parts_relation_def) apply( drule subterm_atom) apply(unfold Atoms_def, auto) done moreover {assume c0:"Crypt ?k ?h \<sqsubset> H0" from c0 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',0)" by (unfold parts_relation_def,simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',0):fst b & ((fst m',0), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 0)" in spec) apply blast done from this have ?thesis by blast } moreover {assume c0:"(B0= B & M0=M &A0=A& N0=Nb)" from a12 and a91 and a92 have "(fst m',1):fst b" apply - apply(subgoal_tac "(fst m',1)=m'") apply(unfold strand_def index_def nodes_def) apply auto done from c0 and a16 and this have ?thesis by blast } ultimately have ?thesis by blast } moreover { assume a91:"snd m'=3" from a91 have a92:"m'=(fst m',3)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and a92 have b10:" Crypt ?k ?h \<sqsubset> H0'" apply - apply(unfold parts_relation_def, auto) apply(subgoal_tac "m'=(fst m',3)") apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(erule disjE) apply simp apply(fold parts_relation_def) apply( drule subterm_atom) apply(unfold Atoms_def, auto) done from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',2)" by (unfold parts_relation_def, simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="2" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',2):fst b & ((fst m',2), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 2)" in spec) apply blast done from this have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a15:" (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0)" from this obtain A0 and B0 and Na0 and Nb0 and M0 and K0 where a16:" is_server (fst m') A0 B0 Na0 Nb0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have "snd m'=1" apply - apply(unfold is_server_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0,Crypt (shrK A0) {|Nonce Na0, Key K0|}, Crypt (shrK B0) {|Nonce Nb0,Key K0|}|}&M0:T" by(unfold is_server_def, unfold node_term_def, auto) from a12 and this have "False " apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this and a16 have ?thesis by blast } ultimately show ?thesis by blast qed from this show ?thesis by(unfold nodes_def,blast) qed lemma server_auth_resp1: assumes a1:"b: bundles" and a2:" is_server s A B Na Nb M K" and a3:" (s,1): nodes b" and a5:"~B:bad" and a6:"A=B" shows " (EX s Ks. is_initiator s A B Nb M Ks & (s,0):nodes b )| (EX r K H H'. is_responder r A B Nb M K H H' & (r,1):nodes b ) " proof - from a2 have a7: "node_term (s,0)= {|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|}, Crypt (shrK B) {|Nonce Nb, M, Agent A, Agent B|} |}& node_term (s,1)={|M, Crypt (shrK A) {|Nonce Na, Key K|}, Crypt (shrK B) {|Nonce Nb,Key K|}|} " apply - apply(unfold is_server_def,unfold node_term_def, auto) done from a1 and a3 have a8:"(s,0): nodes b & ((s,0), (s,1)): (edges b)^*" apply - apply(drule_tac i="0" and ?n2.0="(s,1)" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this have a8:"(s,0): nodes b & ((s,0),(s,1)): (edges b )^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done let ?n="(s,0)" let ?k="shrK B" let ?h="{|Nonce Nb, M, Agent A, Agent B|}" from a1 and a5 have a9:"ALL n:(nodes b). node_term n = Key ?k -->Is_regular_strand (strand n)" proof(rule_tac A="nodes b" in ballI) fix x assume a10:"b:bundles" and a11:"x:nodes b" and a11a:"~B:bad" show " node_term x = Key (shrK B) -->Is_regular_strand (strand x) " proof assume a12:" node_term x = Key (shrK B)" show "Is_regular_strand (strand x)" proof - from a11a and a10 and a11 have a13:" ~node_term x=Key (shrK B)" by(blast dest:no_node_is_key) from a12 and a13 show ?thesis by auto qed qed qed from a7 have a10:"Crypt ?k ?h \<sqsubset> node_term ?n" apply - apply(unfold parts_relation_def, simp add:parts_insert2) done from a1 and a8 and a9 and a10 have a11:"EX n'. (n',?n):(edges b)^*&(n': (nodes b) & Crypt ?k ?h \<sqsubset> node_term n' & Is_regular_strand (strand n') & node_sign n'=+& (ALL y.(y,n'):(edges b)^+ -->~Crypt ?k ?h \<sqsubset> node_term y))" (is "EX n'.?P n'") apply - apply(rule unsolicited_test) apply auto apply(unfold regularK_def) apply assumption done from a11 obtain m' where a12:"?P m'" by blast from this have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') | (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" apply - apply(subgoal_tac "(strand m'):Σ") thm bundle_nodes_strand_in_space prefer 2 apply(fold strand_def) apply(cut_tac a1 ) apply(blast intro: bundle_nodes_strand_in_space) apply(unfold OR_1 Is_regular_strand_def strand_def) apply auto done show ?thesis proof - from a13 have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') |(EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" by simp moreover {assume a15:"(EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0)" from this obtain A0 and B0 and Na0 and M0 and K0 where a16:" is_initiator (fst m') A0 B0 Na0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have b17:"snd m'=0" apply - apply(unfold is_initiator_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def strand_def index_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|}&M0:T" by(unfold is_initiator_def, unfold node_term_def, auto) from a12 and this have "A0=A & B0=B& Na0=Nb & M0=M &A0=B" apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(erule disjE,blast dest: subterm_atom) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this and a16 and b17 and a12 have ?thesis apply - apply(rule disjI1) apply(rule_tac x="fst m'" in exI) apply(rule_tac x="K0" in exI) apply (simp,subgoal_tac "m'=(fst m',0)") apply(erule conjE, simp ) apply(simp add:node_equal) apply(unfold strand_def index_def) apply auto done } moreover {assume a15:"(EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')" from this obtain A0 and B0 and N0 and M0 and K0 and H0 and H0' where a16:" is_responder (fst m') A0 B0 N0 M0 K0 H0 H0'" by blast from a16 have a17: "node_term (fst m',0)={|M0, Agent A0, Agent B0, H0|}& node_term (fst m',1)={|M0, Agent A0, Agent B0, H0, Crypt (shrK B0) {|Nonce N0, M0, Agent A0, Agent B0|} |}& node_term (fst m',2)= {|M0, H0', Crypt (shrK B0) {|Nonce N0,Key K0|}|} & node_term (fst m',3)= {|M0,H0'|}& M0:T" apply - apply(unfold is_responder_def, unfold node_term_def, auto) done from a1 and a12 have a18:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and a16 and a18 have a9:"snd m'=1 | snd m'=3" apply - apply(unfold is_responder_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply simp apply(case_tac "snd m'=3") apply simp apply(case_tac "snd m'=2") apply auto done moreover { assume a91:"snd m'=1" from a91 have a92:"m'=(fst m',1)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0 | (B0= B & M0=M &A0=A& N0=Nb) " apply - apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply((erule conjE)+,erule disjE) apply simp+ apply(rule disjI2) apply(subgoal_tac "A0=B") apply simp apply(erule conjE)+ apply simp apply(erule disjE) apply blast apply(fold parts_relation_def) apply(unfold Atoms_def,simp) apply(unfold parts_relation_def) apply auto done moreover {assume c0:"Crypt ?k ?h \<sqsubset> H0" from c0 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',0)" by (unfold parts_relation_def,simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',0):fst b & ((fst m',0), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 0)" in spec) apply blast done from this have ?thesis by blast } moreover {assume c0:"(B0= B & M0=M &A0=A& N0=Nb)" from a12 and a91 and a92 have "(fst m',1):fst b" apply - apply(subgoal_tac "(fst m',1)=m'") apply(unfold strand_def index_def nodes_def) apply auto done from c0 and a16 and this have ?thesis apply - apply(fold nodes_def) by blast } ultimately have ?thesis by blast } moreover { assume a91:"snd m'=3" from a91 have a92:"m'=(fst m',3)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and a92 have b10:" Crypt ?k ?h \<sqsubset> H0'" apply - apply(unfold parts_relation_def, auto) apply(subgoal_tac "m'=(fst m',3)") apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(erule disjE) apply simp apply(fold parts_relation_def) apply( drule subterm_atom) apply(unfold Atoms_def, auto) done from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',2)" by (unfold parts_relation_def, simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="2" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',2):fst b & ((fst m',2), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 2)" in spec) apply blast done from this have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a15:" (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0)" from this obtain A0 and B0 and Na0 and Nb0 and M0 and K0 where a16:" is_server (fst m') A0 B0 Na0 Nb0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have "snd m'=1" apply - apply(unfold is_server_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0,Crypt (shrK A0) {|Nonce Na0, Key K0|}, Crypt (shrK B0) {|Nonce Nb0,Key K0|}|}&M0:T" by(unfold is_server_def, unfold node_term_def, auto) from a12 and this have "False " apply - apply simp apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(fold parts_relation_def) apply(simp,drule subterm_atom) apply simp apply(unfold Atoms_def,auto) done from this and a16 have ?thesis by blast } ultimately show ?thesis by blast qed qed lemma server_auth_resp_general: assumes a1:"b: bundles" and a2:" is_server s A B Na Nb M K" and a3:" (s,1): nodes b" and a5:"~B:bad" shows " (EX s Ks. is_initiator s A B Nb M Ks & (s,0):nodes b )| (EX r K H H'. is_responder r A B Nb M K H H' & (r,1):nodes b ) " proof - from a1 and a2 and a3 and a5 show ?thesis by (blast dest:server_auth_resp1 server_auth_resp0) qed lemma initiator_auth: assumes a1:"b: bundles" and a2:" is_initiator s A B Na M K" and a3:" (s,1): nodes b" and a4:"~A:bad" and a5:"unique_originate (Nonce Na) (s,0)" and a6:"~A=B" shows " (EX s Nb . is_server s A B Na Nb M K & (s,1):nodes b ) " proof - from a2 have a7: "node_term (s,0)= {|M, Agent A, Agent B, Crypt (shrK A) {|Nonce Na, M, Agent A, Agent B|} |}& node_term (s,1)={|M, Crypt (shrK A) {|Nonce Na, Key K|}|}&M:T " apply - apply(unfold is_initiator_def, unfold node_term_def, auto) done from a1 and a3 have a8:"(s,0): nodes b & ((s,0), (s,1)): (edges b)^*" apply - apply(drule_tac i="0" and ?n2.0="(s,1)" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this have a8:"(s,0): nodes b & ((s,0),(s,1)): (edges b )^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done let ?n="(s,1)" let ?k="shrK A" let ?h="{|Nonce Na, Key K|}" from a1 and a4 have a9:"ALL n:(nodes b). node_term n = Key ?k -->Is_regular_strand (strand n)" proof(rule_tac A="nodes b" in ballI) fix x assume a10:"b:bundles" and a11:"x:nodes b" and a11a:"~A:bad" show " node_term x = Key ?k -->Is_regular_strand (strand x) " proof assume a12:" node_term x = Key ?k" show "Is_regular_strand (strand x)" proof - from a11a and a10 and a11 have a13:" ~node_term x=Key (?k)" by(blast dest:no_node_is_key) from a12 and a13 show ?thesis by auto qed qed qed from a7 have a10:"Crypt ?k ?h \<sqsubset> node_term ?n" apply - apply(unfold parts_relation_def, simp add:parts_insert2) done from a1 and a8 and a9 and a10 and a3 have a11:"EX n'.(n',?n):(edges b)^*& (n': (nodes b) & Crypt ?k ?h \<sqsubset> node_term n' & Is_regular_strand (strand n') & node_sign n'=+& (ALL y.(y,n'):(edges b)^+ -->~Crypt ?k ?h \<sqsubset> node_term y))" (is "EX n'.?P n'") apply - apply(rule unsolicited_test) apply auto apply(unfold regularK_def) apply assumption done from a11 obtain m' where a12:"?P m'" by blast from this have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') | (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" apply - apply(subgoal_tac "(strand m'):Σ") thm bundle_nodes_strand_in_space prefer 2 apply(fold strand_def) apply(cut_tac a1 ) apply(blast intro: bundle_nodes_strand_in_space) apply(unfold OR_1 Is_regular_strand_def strand_def) apply auto done show ?thesis proof - from a13 have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')| (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" by simp moreover {assume a15:"(EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0)" from this obtain A0 and B0 and Na0 and M0 and K0 where a16:" is_initiator (fst m') A0 B0 Na0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have b17:"snd m'=0" apply - apply(unfold is_initiator_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|}&M0:T" by(unfold is_initiator_def, unfold node_term_def, auto) from this and a12 have "Crypt ?k ?h \<sqsubset> {|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|} &M0:T" by simp from this have False apply - apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(unfold Atoms_def,auto) done then have ?thesis by blast } moreover {assume a15:"(EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')" from this obtain A0 and B0 and N0 and M0 and K0 and H0 and H0' where a16:" is_responder (fst m') A0 B0 N0 M0 K0 H0 H0'" by blast from a16 have a17: "node_term (fst m',0)={|M0, Agent A0, Agent B0, H0|}& node_term (fst m',1)={|M0, Agent A0, Agent B0, H0, Crypt (shrK B0) {|Nonce N0, M0, Agent A0, Agent B0|} |}& node_term (fst m',2)= {|M0, H0', Crypt (shrK B0) {|Nonce N0,Key K0|}|} & node_term (fst m',3)= {|M0,H0'|}& M0:T" apply - apply(unfold is_responder_def, unfold node_term_def, auto) done from a1 and a12 have a18:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and a16 and a18 have a9:"snd m'=1 | snd m'=3" apply - apply(unfold is_responder_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply simp apply(case_tac "snd m'=3") apply simp apply(case_tac "snd m'=2") apply auto done moreover { assume a91:"snd m'=1" from a91 have a92:"m'=(fst m',1)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0 " apply - apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply((erule conjE)+,erule disjE) apply simp+ apply(unfold Atoms_def, auto) done from this and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',0)" by (unfold parts_relation_def,simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',0):fst b & ((fst m',0), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 0)" in spec) apply blast done from this have ?thesis by blast } moreover { assume a91:"snd m'=3" from a91 have a92:"m'=(fst m',3)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0'" apply - apply(unfold parts_relation_def, auto) apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(erule disjE) apply simp apply(fold parts_relation_def) apply( drule subterm_atom) apply(unfold Atoms_def, auto) done from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',2)" by (unfold parts_relation_def, simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="2" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',2):fst b & ((fst m',2), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 2)" in spec) apply blast done then have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a15:" (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0)" from this obtain A0 and B0 and Na0 and Nb0 and M0 and K0 where a16:" is_server (fst m') A0 B0 Na0 Nb0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have a17a:"snd m'=1" apply - apply(unfold is_server_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0,Crypt (shrK A0) {|Nonce Na0, Key K0|}, Crypt (shrK B0) {|Nonce Nb0,Key K0|}|}&M0:T" by(unfold is_server_def, unfold node_term_def, auto) from a12 and this have a19:"Na0=Na&A0=A&K0=K|Nb0=Na&B0=A&K0=K" apply - apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(unfold Atoms_def,auto) done moreover {assume a20: "Na0=Na&A0=A&K0=K" from a17a have a20a:"m'=(fst m',1)" by (simp add:node_equal, unfold strand_def index_def, auto) from a1 and a12 and a16 and a4 and this and a20 have a21:"(EX s Ks. is_initiator s A0 B0 Na M0 Ks & (s,0):fst b ) |(EX r K H H'. is_responder r A0 B0 Na M0 K H H'&(r,1):fst b ) " apply - apply(fold nodes_def) apply (drule server_auth_init) apply auto done moreover {assume a21:"EX s Ks. is_initiator s A0 B0 Na M0 Ks" from a21 obtain s0 Ks where a22:"is_initiator s0 A0 B0 Na M0 Ks" by auto from a2 and this and a5 have "s=s0&A=A0&B=B0&M=M0&Ks=K" apply - thm unique_orig_init apply (drule unique_orig_init,auto) done from this and a19 and a16 and a17a and a12 and a20a and a20 have ?thesis apply - apply(rule_tac x="fst m'" in exI,rule_tac x="Nb0" in exI ) apply auto done } moreover {assume a21:"(EX r K H H'. is_responder r A0 B0 Na M0 K H H'&(r,1):fst b )" from a21 obtain r K1 H1 H1' where a22:"is_responder r A0 B0 Na M0 K1 H1 H1'" by blast from a20 this have a20a:"~Nonce Na \<sqsubset> node_term (r,0)" apply - apply(unfold originate_def is_responder_def T_is_only_text Atoms_def node_sign_def node_term_def strand_def parts_relation_def ) apply( simp add:parts_insert2) apply auto done from a22 have a20b:"Nonce Na \<sqsubset> node_term (r,1) &node_sign (r,1)=+" apply - apply(unfold originate_def is_responder_def T_is_only_text Atoms_def node_sign_def node_term_def strand_def parts_relation_def ) apply( simp add:parts_insert2) done from a20a a20b have a23:"originate (Nonce Na) (r,1)" apply - apply(unfold originate_def index_def strand_def) apply(simp add: parts_insert2) done from this and a5 have "s=r" apply - apply(unfold unique_originate_def) apply auto done with a22 a2 have False apply - apply(unfold is_initiator_def, unfold is_responder_def) apply auto done then have ?thesis by blast } ultimately have ?thesis by blast } moreover { assume a20:"Nb0=Na&B0=A&K0=K" from a17a have a20a:"m'=(fst m',1)" by ( simp add:node_equal, unfold strand_def index_def, auto) from a1 and a12 and a16 and a4 and this and a20 have a21:"(EX s Ks. is_initiator s A0 B0 Na M0 Ks & (s,0):nodes b)| (EX r K H H'. is_responder r A0 B0 Na M0 K H H' & (r,1):nodes b )" apply - apply(drule_tac A="A0" and B="B0" and Nb="Nb0" and M="M0" and Na="Na0" in server_auth_resp_general, auto ) done moreover {assume a21:"EX s Ks. is_initiator s A0 B0 Na M0 Ks" from a21 obtain s0 Ks where a22:"is_initiator s0 A0 B0 Na M0 Ks" by auto from a2 this a5 have "s=s0&A=A0&B=B0&M=M0&Ks=K" apply - apply (drule unique_orig_init,auto) done from this and a20 have "A=B" by simp from this and a6 have ?thesis by simp } moreover {assume a21:"(EX r K H H'. is_responder r A0 B0 Na M0 K H H')" from a21 obtain r K1 H1 H1' where a22:"is_responder r A0 B0 Na M0 K1 H1 H1'" by blast from a20 this have a20a:"~Nonce Na \<sqsubset> node_term (r,0)" apply - apply(unfold originate_def is_responder_def T_is_only_text Atoms_def node_sign_def node_term_def strand_def parts_relation_def ) apply( simp add:parts_insert2) apply auto done from a22 have a20b:"Nonce Na \<sqsubset> node_term (r,1) &node_sign (r,1)=+" apply - apply(unfold originate_def is_responder_def T_is_only_text Atoms_def node_sign_def node_term_def strand_def parts_relation_def ) apply( simp add:parts_insert2) done from a20a a20b have a23:"originate (Nonce Na) (r,1)" apply - apply(unfold originate_def index_def strand_def) apply(simp add: parts_insert2) done from this and a5 have "s=r" apply - apply(unfold unique_originate_def) apply auto done with a22 a2 have False apply - apply(unfold is_initiator_def, unfold is_responder_def) apply auto done then have ?thesis by blast } ultimately have ?thesis by blast } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed qed lemma unique_orig_resp: assumes a1:"is_responder s A B Nb M K H H'" and a2:"is_responder s' A' B' Nb M' K' H0 H0'" and a3:"unique_originate (Nonce Nb) (s,1)" shows "s=s'& A=A'& B=B'&M=M'&K=K'&H=H0&H'=H0'" proof - from a2 have "Nonce Nb \<sqsubset> node_term (s',1)" apply - apply(unfold parts_relation_def is_responder_def node_term_def) apply(simp add:parts_insert2) done from a2 have a2a:"~Nonce Nb \<sqsubset> node_term (s',0)" apply - apply(unfold originate_def is_responder_def T_is_only_text Atoms_def node_sign_def node_term_def strand_def parts_relation_def ) apply( simp add:parts_insert2) apply auto done from a2 have a2b:"Nonce Nb \<sqsubset> node_term (s',1) &node_sign (s',1)=+" apply - apply(unfold originate_def is_responder_def T_is_only_text Atoms_def node_sign_def node_term_def strand_def parts_relation_def ) apply( simp add:parts_insert2) done from this a2a have "originate (Nonce Nb) (s',1)" apply - apply(unfold originate_def index_def strand_def) apply(simp add: parts_insert2) done from this a1 a2 a3 show ?thesis apply - apply(unfold is_responder_def unique_originate_def) apply(erule conjE)+ apply(drule_tac x="(s',1)" in spec) apply simp done qed lemma resp_auth: assumes a1:"b: bundles" and a2:" is_responder s A B Nb M K H H'" and a3:" (s,3):nodes b" and a4:"~B:bad" and a5:"unique_originate (Nonce Nb) (s,1)" and a6:"~A=B" shows " (EX s Na . is_server s A B Na Nb M K & (s,1):nodes b ) " proof - from a2 have a7: "node_term (s,0)={|M, Agent A, Agent B, H|} &node_term (s,1)={|M, Agent A, Agent B, H, Crypt (shrK B) {|Nonce Nb, M, Agent A, Agent B|} |}& node_term (s,2)= {|M, H', Crypt (shrK B) {|Nonce Nb,Key K|}|} & node_term (s,3)= {|M,H'|}& M:T " apply - apply(unfold is_responder_def, unfold node_term_def, auto) done from a1 and a3 have a8:"(s,2): nodes b & ((s,2), (s,3)): (edges b)^*" apply - apply(drule_tac i="2" and ?n2.0="(s,3)" in bundle_casual2_closed3) apply( simp) apply(unfold index_def,simp) apply(unfold strand_def, simp)thm rtrancl_mono1 apply(rule_tac rtrancl_mono1) apply blast+ done from this have a8:"(s,2): nodes b & ((s,2),(s,3)): (edges b )^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done let ?n="(s,2)" let ?k="shrK B" let ?h="{|Nonce Nb, Key K|}" from a1 and a4 have a9:"ALL n:(nodes b). node_term n = Key ?k -->Is_regular_strand (strand n)" proof(rule_tac A="nodes b" in ballI) fix x assume a10:"b:bundles" and a11:"x:nodes b" and a11a:"~B:bad" show " node_term x = Key ?k -->Is_regular_strand (strand x) " proof assume a12:" node_term x = Key ?k" show "Is_regular_strand (strand x)" proof - from a11a and a10 and a11 have a13:" ~node_term x=Key (?k)" by(blast dest:no_node_is_key) from a12 and a13 show ?thesis by auto qed qed qed from a7 have a10:"Crypt ?k ?h \<sqsubset> node_term ?n" apply - apply(unfold parts_relation_def, simp add:parts_insert2) done from a1 and a8 and a9 and a10 and a3 have a11:"EX n'.(n',?n):(edges b)^*& (n': (nodes b) & Crypt ?k ?h \<sqsubset> node_term n' & Is_regular_strand (strand n') & node_sign n'=+& (ALL y.(y,n'):(edges b)^+ -->~Crypt ?k ?h \<sqsubset> node_term y))" (is "EX n'.?P n'") apply - apply(rule unsolicited_test) apply auto apply(unfold regularK_def) apply assumption done from a11 obtain m' where a12:"?P m'" by blast from this have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0') | (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" apply - apply(subgoal_tac "(strand m'):Σ") thm bundle_nodes_strand_in_space prefer 2 apply(fold strand_def) apply(cut_tac a1 ) apply(blast intro: bundle_nodes_strand_in_space) apply(unfold OR_1 Is_regular_strand_def strand_def) apply auto done show ?thesis proof - from a13 have a13:" ((EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0) | (EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')| (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0))" by simp moreover {assume a15:"(EX A0 B0 Na0 M0 K0. is_initiator (fst m') A0 B0 Na0 M0 K0)" from this obtain A0 and B0 and Na0 and M0 and K0 where a16:" is_initiator (fst m') A0 B0 Na0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have b17:"snd m'=0" apply - apply(unfold is_initiator_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18: "node_term m'={|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|}&M0:T" by(unfold is_initiator_def, unfold node_term_def, auto) from this and a12 have "Crypt ?k ?h \<sqsubset> {|M0, Agent A0, Agent B0, Crypt (shrK A0) {|Nonce Na0, M0, Agent A0, Agent B0|}|}&M0:T" by simp from this have False apply - apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(unfold Atoms_def,auto) done then have ?thesis by blast } moreover {assume a15:"(EX A0 B0 Nb0 M0 K0 H0 H0'. is_responder (fst m') A0 B0 Nb0 M0 K0 H0 H0')" from this obtain A0 and B0 and N0 and M0 and K0 and H0 and H0' where a16:" is_responder (fst m') A0 B0 N0 M0 K0 H0 H0'" by blast from a16 have a17: "node_term (fst m',0)={|M0, Agent A0, Agent B0, H0|}& node_term (fst m',1)={|M0, Agent A0, Agent B0, H0, Crypt (shrK B0) {|Nonce N0, M0, Agent A0, Agent B0|} |}& node_term (fst m',2)= {|M0, H0', Crypt (shrK B0) {|Nonce N0,Key K0|}|} & node_term (fst m',3)= {|M0,H0'|}& M0:T" apply - apply(unfold is_responder_def,unfold node_term_def, auto) done from a1 and a12 have a18:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and a16 and a18 have a9:"snd m'=1 | snd m'=3" apply - apply(unfold is_responder_def) apply(unfold node_sign_def node_term_def Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply simp apply(case_tac "snd m'=3") apply simp apply(case_tac "snd m'=2") apply auto done moreover { assume a91:"snd m'=1" from a91 have a92:"m'=(fst m',1)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0 " apply - apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply((erule conjE)+,erule disjE) apply simp+ apply(unfold Atoms_def, auto) done from this and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',0)" by (unfold parts_relation_def,simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',0): fst b & ((fst m',0),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',0):fst b & ((fst m',0), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 0)" in spec) apply blast done from this have ?thesis by blast } moreover { assume a91:"snd m'=3" from a91 have a92:"m'=(fst m',3)" by(simp add:node_equal, unfold strand_def index_def, auto) from a91 and a17 and a12 and a6 and this have b10:" Crypt ?k ?h \<sqsubset> H0'" apply - apply(unfold parts_relation_def, auto) apply (simp add:parts_insert2) apply(unfold T_is_only_text) apply(erule disjE) apply simp apply(fold parts_relation_def) apply( drule subterm_atom) apply(unfold Atoms_def, auto) done from b10 and a17 have b12:"Crypt ?k ?h \<sqsubset> node_term (fst m',2)" by (unfold parts_relation_def, simp add:parts_insert2) from a1 and a91 and a12 have N3:"(fst m',2): fst b & ((fst m',2),m'): (snd b)^*" apply - apply(fold edges_def nodes_def index_def strand_def) apply(drule_tac i="2" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply(unfold strand_def, simp)thm rtrancl_mono1 apply( rule_tac rtrancl_mono1) apply blast+ done from this a91 have N3:"(fst m',2):fst b & ((fst m',2), m'): (snd b)^+" apply - apply(erule conjE,drule rtranclD) apply(erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done from b12 and a12 and N3 have False apply- apply(erule conjE)+ apply(fold edges_def nodes_def) apply(drule_tac x="(fst m', 2)" in spec) apply blast done then have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a15:" (EX A0 B0 Na0 Nb0 M0 K0. is_server (fst m') A0 B0 Na0 Nb0 M0 K0)" from this obtain A0 and B0 and Na0 and Nb0 and M0 and K0 where a16:" is_server (fst m') A0 B0 Na0 Nb0 M0 K0" by blast from a1 and a12 have a17:" m': Domain" by(blast dest:bundle_node_in_Domain) from a12 and this and a16 have a17a:"snd m'=1" apply - apply(unfold is_server_def) apply(unfold node_sign_def, unfold node_term_def, unfold Domain_def) apply(unfold parts_relation_def) apply(case_tac "snd m'=0") apply simp apply(case_tac "snd m'=1") apply auto done from this and a16 have a18:"node_term m'={|M0,Crypt (shrK A0) {|Nonce Na0, Key K0|}, Crypt (shrK B0) {|Nonce Nb0,Key K0|}|}&M0:T" by(unfold is_server_def, unfold node_term_def, auto) from a12 and this have a19:"Na0=Nb&A0=B&K0=K|Nb0=Nb&B0=B&K0=K" apply - apply(erule conjE)+ apply(unfold parts_relation_def) apply (simp add:parts_insert2)+ apply(unfold T_is_only_text) apply(unfold Atoms_def,auto) done moreover {assume a20: "Na0=Nb&A0=B&K0=K" from a17a have a20a:"m'=(fst m',1)" by(simp add:node_equal, unfold strand_def index_def, auto) from a1 and a12 and a16 and a4 and this and a20 have a21:"(EX s Ks. is_initiator s A0 B0 Na0 M0 Ks & (s,0):fst b ) | (EX r K H H'. is_responder r A0 B0 Na0 M0 K H H'&(r,1):fst b ) " apply - apply(fold nodes_def) apply (drule server_auth_init) apply auto done moreover {assume a21:"EX s Ks. is_initiator s A0 B0 Na0 M0 Ks" from a21 obtain s0 Ks where a22:"is_initiator s0 A0 B0 Na0 M0 Ks" by auto from a20 a22 have a23:"originate (Nonce Nb) (s0,0)" apply - apply(unfold is_initiator_def originate_def parts_relation_def index_def strand_def node_term_def node_sign_def) apply(simp add: parts_insert2) done from this and a5 have "s=s0" apply - apply(unfold unique_originate_def) apply auto done with a22 a2 have False apply - apply(unfold is_initiator_def, unfold is_responder_def) apply auto done then have ?thesis by blast } moreover {assume a21:"(EX r K H H'. is_responder r A0 B0 Na0 M0 K H H'&(r,1):fst b )" from a21 obtain r K1 H1 H1' where a22:"is_responder r A0 B0 Na0 M0 K1 H1 H1'" by blast from this and a20 have a23:"is_responder r A0 B0 Nb M0 K1 H1 H1'" by simp from a2 a5 this have a24:"A0=A & B0=B&M0=M&K1=K" by (blast dest:unique_orig_resp) from this and a20 have "A=B" by simp from this and a6 have False by simp then have ?thesis by simp } ultimately have ?thesis by blast } moreover { assume a20:"Nb0=Nb&B0=B&K0=K" from a17a have a20a:"m'=(fst m',1)" by(simp add:node_equal, unfold strand_def index_def, auto) from a1 and a12 and a16 and a4 and this and a20 have a21:"(EX s Ks. is_initiator s A0 B0 Nb M0 Ks & (s,0):nodes b )| (EX r K H H'. is_responder r A0 B0 Nb M0 K H H' & (r,1):nodes b )" apply - apply(drule_tac A="A0" and B="B0" and Nb="Nb0" and M="M0" and Na="Na0" in server_auth_resp_general, auto ) done moreover {assume a21:"EX s Ks. is_initiator s A0 B0 Nb M0 Ks" from a21 obtain s0 Ks where a22:"is_initiator s0 A0 B0 Nb M0 Ks" by auto from a22 have a23:"originate (Nonce Nb) (s0,0)" apply - apply(unfold is_initiator_def originate_def parts_relation_def index_def strand_def node_term_def node_sign_def) apply(simp add: parts_insert2) done from this and a5 have "s=s0" apply - apply(unfold unique_originate_def) apply auto done with a22 a2 have False apply - apply(unfold is_initiator_def, unfold is_responder_def) apply auto done then have ?thesis by blast } moreover {assume a21:"(EX r K H H'. is_responder r A0 B0 Nb M0 K H H'&(r,1):fst b )" from a21 obtain r K1 H1 H1' where a22:"is_responder r A0 B0 Nb M0 K1 H1 H1'" by blast from a2 this a5 have a24:"A0=A & B0=B&M0=M&K1=K" by (blast dest:unique_orig_resp) from this and a20 and a16 and a20a and a12 have ?thesis apply - apply (rule_tac x="fst m'" in exI, rule_tac x="Na0" in exI, auto) done } ultimately have ?thesis by(fold nodes_def, blast) } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed qed end
lemma c:
[|b ∈ bundles; n ∈ nodes b|] ==> ∀ag. ag ∉ bad --> node_term n ≠ Key (shrK ag)
lemma server_auth:
[|b ∈ bundles; is_server s A B Na Nb M K; (s, 1) ∈ nodes b; A ∉ bad; A ≠ B|] ==> ∃s Ks. is_initiator s A B Na M Ks ∧ (s, 0) ∈ nodes b
lemma server_auth1:
[|b ∈ bundles; is_server s A B Na Nb M K; (s, 1) ∈ nodes b; A ∉ bad; A = B|] ==> (∃s Ks. is_initiator s A B Na M Ks ∧ (s, 0) ∈ nodes b) ∨ (∃r K H H'. is_responder r A B Na M K H H' ∧ (r, 1) ∈ nodes b)
lemma server_auth_init:
[|b ∈ bundles; is_server s A B Na Nb M K; (s, 1) ∈ nodes b; A ∉ bad|] ==> (∃s Ks. is_initiator s A B Na M Ks ∧ (s, 0) ∈ nodes b) ∨ (∃r K H H'. is_responder r A B Na M K H H' ∧ (r, 1) ∈ nodes b)
lemma unique_orig_init:
[|is_initiator s A B Na M K; is_initiator s' A' B' Na M' K'; unique_originate (Nonce Na) (s, 0)|] ==> s = s' ∧ A = A' ∧ B = B' ∧ M = M' ∧ K = K'
lemma server_auth_resp0:
[|b ∈ bundles; is_server s A B Na Nb M K; (s, 1) ∈ nodes b; B ∉ bad; A ≠ B|] ==> ∃r K H H'. is_responder r A B Nb M K H H' ∧ (r, 1) ∈ nodes b
lemma server_auth_resp1:
[|b ∈ bundles; is_server s A B Na Nb M K; (s, 1) ∈ nodes b; B ∉ bad; A = B|] ==> (∃s Ks. is_initiator s A B Nb M Ks ∧ (s, 0) ∈ nodes b) ∨ (∃r K H H'. is_responder r A B Nb M K H H' ∧ (r, 1) ∈ nodes b)
lemma server_auth_resp_general:
[|b ∈ bundles; is_server s A B Na Nb M K; (s, 1) ∈ nodes b; B ∉ bad|] ==> (∃s Ks. is_initiator s A B Nb M Ks ∧ (s, 0) ∈ nodes b) ∨ (∃r K H H'. is_responder r A B Nb M K H H' ∧ (r, 1) ∈ nodes b)
lemma initiator_auth:
[|b ∈ bundles; is_initiator s A B Na M K; (s, 1) ∈ nodes b; A ∉ bad; unique_originate (Nonce Na) (s, 0); A ≠ B|] ==> ∃s Nb. is_server s A B Na Nb M K ∧ (s, 1) ∈ nodes b
lemma unique_orig_resp:
[|is_responder s A B Nb M K H H'; is_responder s' A' B' Nb M' K' H0.0 H0'; unique_originate (Nonce Nb) (s, 1)|] ==> s = s' ∧ A = A' ∧ B = B' ∧ M = M' ∧ K = K' ∧ H = H0.0 ∧ H' = H0'
lemma resp_auth:
[|b ∈ bundles; is_responder s A B Nb M K H H'; (s, 3) ∈ nodes b; B ∉ bad; unique_originate (Nonce Nb) (s, 1); A ≠ B|] ==> ∃s Na. is_server s A B Na Nb M K ∧ (s, 1) ∈ nodes b