| author | wenzelm | 
| Mon, 05 May 2014 20:10:33 +0200 | |
| changeset 56871 | d06ff36b4fa7 | 
| parent 56073 | 29e308b56d23 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 41775 | 1 | (* Title: HOL/Auth/Guard/Guard_Yahalom.thy | 
| 2 | Author: Frederic Blanqui, University of Cambridge Computer Laboratory | |
| 3 | Copyright 2002 University of Cambridge | |
| 4 | *) | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 5 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 6 | header{*Yahalom Protocol*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 7 | |
| 39216 | 8 | theory Guard_Yahalom imports "../Shared" Guard_Shared begin | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 9 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 10 | subsection{*messages used in the protocol*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 11 | |
| 20768 | 12 | abbreviation (input) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 13 | ya1 :: "agent => agent => nat => event" where | 
| 20768 | 14 |   "ya1 A B NA == Says A B {|Agent A, Nonce NA|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 15 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 16 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 17 | ya1' :: "agent => agent => agent => nat => event" where | 
| 20768 | 18 |   "ya1' A' A B NA == Says A' B {|Agent A, Nonce NA|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 19 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 20 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 21 | ya2 :: "agent => agent => nat => nat => event" where | 
| 20768 | 22 |   "ya2 A B NA NB == Says B Server {|Agent B, Ciph B {|Agent A, Nonce NA, Nonce NB|}|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 23 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 24 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 25 | ya2' :: "agent => agent => agent => nat => nat => event" where | 
| 20768 | 26 |   "ya2' B' A B NA NB == Says B' Server {|Agent B, Ciph B {|Agent A, Nonce NA, Nonce NB|}|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 27 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 28 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 29 | ya3 :: "agent => agent => nat => nat => key => event" where | 
| 20768 | 30 | "ya3 A B NA NB K == | 
| 31 |     Says Server A {|Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|},
 | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 32 |                     Ciph B {|Agent A, Key K|}|}"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 33 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 34 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 35 | ya3':: "agent => msg => agent => agent => nat => nat => key => event" where | 
| 20768 | 36 | "ya3' S Y A B NA NB K == | 
| 37 |     Says S A {|Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|}, Y|}"
 | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 38 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 39 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 40 | ya4 :: "agent => agent => nat => nat => msg => event" where | 
| 20768 | 41 |   "ya4 A B K NB Y == Says A B {|Y, Crypt K (Nonce NB)|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 42 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 43 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 44 | ya4' :: "agent => agent => nat => nat => msg => event" where | 
| 20768 | 45 |   "ya4' A' B K NB Y == Says A' B {|Y, Crypt K (Nonce NB)|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 46 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 47 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 48 | subsection{*definition of the protocol*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 49 | |
| 23746 | 50 | inductive_set ya :: "event list set" | 
| 51 | where | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 52 | |
| 23746 | 53 | Nil: "[]:ya" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 54 | |
| 23746 | 55 | | Fake: "[| evs:ya; X:synth (analz (spies evs)) |] ==> Says Spy B X # evs:ya" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 56 | |
| 23746 | 57 | | YA1: "[| evs1:ya; Nonce NA ~:used evs1 |] ==> ya1 A B NA # evs1:ya" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 58 | |
| 23746 | 59 | | YA2: "[| evs2:ya; ya1' A' A B NA:set evs2; Nonce NB ~:used evs2 |] | 
| 60 | ==> ya2 A B NA NB # evs2:ya" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 61 | |
| 23746 | 62 | | YA3: "[| evs3:ya; ya2' B' A B NA NB:set evs3; Key K ~:used evs3 |] | 
| 63 | ==> ya3 A B NA NB K # evs3:ya" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 64 | |
| 23746 | 65 | | YA4: "[| evs4:ya; ya1 A B NA:set evs4; ya3' S Y A B NA NB K:set evs4 |] | 
| 66 | ==> ya4 A B K NB Y # evs4:ya" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 67 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 68 | subsection{*declarations for tactics*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 69 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 70 | declare knows_Spy_partsEs [elim] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 71 | declare Fake_parts_insert [THEN subsetD, dest] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 72 | declare initState.simps [simp del] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 73 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 74 | subsection{*general properties of ya*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 75 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 76 | lemma ya_has_no_Gets: "evs:ya ==> ALL A X. Gets A X ~:set evs" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 77 | by (erule ya.induct, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 78 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 79 | lemma ya_is_Gets_correct [iff]: "Gets_correct ya" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 80 | by (auto simp: Gets_correct_def dest: ya_has_no_Gets) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 81 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 82 | lemma ya_is_one_step [iff]: "one_step ya" | 
| 23746 | 83 | by (unfold one_step_def, clarify, ind_cases "ev#evs:ya" for ev evs, auto) | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 84 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 85 | lemma ya_has_only_Says' [rule_format]: "evs:ya ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 86 | ev:set evs --> (EX A B X. ev=Says A B X)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 87 | by (erule ya.induct, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 88 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 89 | lemma ya_has_only_Says [iff]: "has_only_Says ya" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 90 | by (auto simp: has_only_Says_def dest: ya_has_only_Says') | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 91 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 92 | lemma ya_is_regular [iff]: "regular ya" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 93 | apply (simp only: regular_def, clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 94 | apply (erule ya.induct, simp_all add: initState.simps knows.simps) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 95 | by (auto dest: parts_sub) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 96 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 97 | subsection{*guardedness of KAB*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 98 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 99 | lemma Guard_KAB [rule_format]: "[| evs:ya; A ~:bad; B ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 100 | ya3 A B NA NB K:set evs --> GuardK K {shrK A,shrK B} (spies evs)" 
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 101 | apply (erule ya.induct) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 102 | (* Nil *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 103 | apply simp_all | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 104 | (* Fake *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 105 | apply (clarify, erule in_synth_GuardK, erule GuardK_analz, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 106 | (* YA1 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 107 | (* YA2 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 108 | apply safe | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 109 | apply (blast dest: Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 110 | (* YA3 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 111 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 112 | apply (drule_tac A=Server in Key_neq, simp+, rule No_Key, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 113 | apply (drule_tac A=Server in Key_neq, simp+, rule No_Key, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 114 | (* YA4 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 115 | apply (blast dest: Says_imp_spies in_GuardK_kparts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 116 | by blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 117 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 118 | subsection{*session keys are not symmetric keys*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 119 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 120 | lemma KAB_isnt_shrK [rule_format]: "evs:ya ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 121 | ya3 A B NA NB K:set evs --> K ~:range shrK" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 122 | by (erule ya.induct, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 123 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 124 | lemma ya3_shrK: "evs:ya ==> ya3 A B NA NB (shrK C) ~:set evs" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 125 | by (blast dest: KAB_isnt_shrK) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 126 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 127 | subsection{*ya2' implies ya1'*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 128 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 129 | lemma ya2'_parts_imp_ya1'_parts [rule_format]: | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 130 | "[| evs:ya; B ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 131 |       Ciph B {|Agent A, Nonce NA, Nonce NB|}:parts (spies evs) -->
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 132 |       {|Agent A, Nonce NA|}:spies evs"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 133 | by (erule ya.induct, auto dest: Says_imp_spies intro: parts_parts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 134 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 135 | lemma ya2'_imp_ya1'_parts: "[| ya2' B' A B NA NB:set evs; evs:ya; B ~:bad |] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 136 | ==> {|Agent A, Nonce NA|}:spies evs"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 137 | by (blast dest: Says_imp_spies ya2'_parts_imp_ya1'_parts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 138 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 139 | subsection{*uniqueness of NB*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 140 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 141 | lemma NB_is_uniq_in_ya2'_parts [rule_format]: "[| evs:ya; B ~:bad; B' ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 142 | Ciph B {|Agent A, Nonce NA, Nonce NB|}:parts (spies evs) -->
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 143 | Ciph B' {|Agent A', Nonce NA', Nonce NB|}:parts (spies evs) -->
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 144 | A=A' & B=B' & NA=NA'" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 145 | apply (erule ya.induct, simp_all, clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 146 | apply (drule Crypt_synth_insert, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 147 | apply (drule Crypt_synth_insert, simp+, safe) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 148 | apply (drule not_used_parts_false, simp+)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 149 | by (drule Says_not_parts, simp+)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 150 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 151 | lemma NB_is_uniq_in_ya2': "[| ya2' C A B NA NB:set evs; | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 152 | ya2' C' A' B' NA' NB:set evs; evs:ya; B ~:bad; B' ~:bad |] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 153 | ==> A=A' & B=B' & NA=NA'" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 154 | by (drule NB_is_uniq_in_ya2'_parts, auto dest: Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 155 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 156 | subsection{*ya3' implies ya2'*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 157 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 158 | lemma ya3'_parts_imp_ya2'_parts [rule_format]: "[| evs:ya; A ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 159 | Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 160 | --> Ciph B {|Agent A, Nonce NA, Nonce NB|}:parts (spies evs)"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 161 | apply (erule ya.induct, simp_all) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 162 | apply (clarify, drule Crypt_synth_insert, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 163 | apply (blast intro: parts_sub, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 164 | by (auto dest: Says_imp_spies parts_parts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 165 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 166 | lemma ya3'_parts_imp_ya2' [rule_format]: "[| evs:ya; A ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 167 | Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 168 | --> (EX B'. ya2' B' A B NA NB:set evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 169 | apply (erule ya.induct, simp_all, safe) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 170 | apply (drule Crypt_synth_insert, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 171 | apply (drule Crypt_synth_insert, simp+, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 172 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 173 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 174 | by (auto dest: Says_imp_spies2 parts_parts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 175 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 176 | lemma ya3'_imp_ya2': "[| ya3' S Y A B NA NB K:set evs; evs:ya; A ~:bad |] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 177 | ==> (EX B'. ya2' B' A B NA NB:set evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 178 | by (drule ya3'_parts_imp_ya2', auto dest: Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 179 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 180 | subsection{*ya3' implies ya3*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 181 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 182 | lemma ya3'_parts_imp_ya3 [rule_format]: "[| evs:ya; A ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 183 | Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|}:parts(spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 184 | --> ya3 A B NA NB K:set evs" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 185 | apply (erule ya.induct, simp_all, safe) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 186 | apply (drule Crypt_synth_insert, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 187 | by (blast dest: Says_imp_spies2 parts_parts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 188 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 189 | lemma ya3'_imp_ya3: "[| ya3' S Y A B NA NB K:set evs; evs:ya; A ~:bad |] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 190 | ==> ya3 A B NA NB K:set evs" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 191 | by (blast dest: Says_imp_spies ya3'_parts_imp_ya3) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 192 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 193 | subsection{*guardedness of NB*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 194 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
23746diff
changeset | 195 | definition ya_keys :: "agent => agent => nat => nat => event list => key set" where | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 196 | "ya_keys A B NA NB evs == {shrK A,shrK B} Un {K. ya3 A B NA NB K:set evs}"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 197 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 198 | lemma Guard_NB [rule_format]: "[| evs:ya; A ~:bad; B ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 199 | ya2 A B NA NB:set evs --> Guard NB (ya_keys A B NA NB evs) (spies evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 200 | apply (erule ya.induct) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 201 | (* Nil *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 202 | apply (simp_all add: ya_keys_def) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 203 | (* Fake *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 204 | apply safe | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 205 | apply (erule in_synth_Guard, erule Guard_analz, simp, clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 206 | apply (frule_tac B=B in Guard_KAB, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 207 | apply (drule_tac p=ya in GuardK_Key_analz, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 208 | apply (blast dest: KAB_isnt_shrK, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 209 | (* YA1 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 210 | apply (drule_tac n=NB in Nonce_neq, simp+, rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 211 | (* YA2 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 212 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 213 | apply (drule Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 214 | apply (drule_tac n=NB in Nonce_neq, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 215 | apply (drule_tac n'=NAa in in_Guard_kparts_neq, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 216 | apply (rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 217 | (* YA3 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 218 | apply (rule Guard_extand, simp, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 219 | apply (case_tac "NAa=NB", clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 220 | apply (frule Says_imp_spies) | 
| 56073 
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
 nipkow parents: 
41775diff
changeset | 221 | apply (frule in_Guard_kparts_Crypt, simp+) | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 222 | apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Ba in ya3_shrK, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 223 | apply (drule ya2'_imp_ya1'_parts, simp, blast, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 224 | apply (case_tac "NBa=NB", clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 225 | apply (frule Says_imp_spies) | 
| 56073 
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
 nipkow parents: 
41775diff
changeset | 226 | apply (frule in_Guard_kparts_Crypt, simp+) | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 227 | apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Ba in ya3_shrK, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 228 | apply (drule NB_is_uniq_in_ya2', simp+, blast, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 229 | apply (simp add: No_Nonce, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 230 | (* YA4 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 231 | apply (blast dest: Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 232 | apply (case_tac "NBa=NB", clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 233 | apply (frule_tac A=S in Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 234 | apply (frule in_Guard_kparts_Crypt, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 235 | apply (blast dest: Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 236 | apply (case_tac "NBa=NB", clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 237 | apply (frule_tac A=S in Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 238 | apply (frule in_Guard_kparts_Crypt, simp+, blast, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 239 | apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Aa in ya3_shrK, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 240 | apply (frule ya3'_imp_ya2', simp+, blast, clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 241 | apply (frule_tac A=B' in Says_imp_spies) | 
| 56073 
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
 nipkow parents: 
41775diff
changeset | 242 | apply (rotate_tac -1, frule in_Guard_kparts_Crypt, simp+) | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 243 | apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Ba in ya3_shrK, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 244 | apply (drule NB_is_uniq_in_ya2', simp+, blast, clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 245 | apply (drule ya3'_imp_ya3, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 246 | apply (simp add: Guard_Nonce) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 247 | apply (simp add: No_Nonce) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 248 | done | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 249 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 250 | end |