| author | huffman | 
| Thu, 17 Nov 2011 14:52:05 +0100 | |
| changeset 45546 | 6dd3e88de4c2 | 
| parent 41775 | 6214816d79d3 | 
| child 48261 | 865610355ef3 | 
| permissions | -rw-r--r-- | 
| 41775 | 1 | (* Title: HOL/Auth/Guard/Guard_NS_Public.thy | 
| 2 | Author: Frederic Blanqui, University of Cambridge Computer Laboratory | |
| 3 | Copyright 2002 University of Cambridge | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 4 | |
| 41775 | 5 | Incorporating Lowe's fix (inclusion of B's identity in round 2). | 
| 6 | *) | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 7 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 8 | header{*Needham-Schroeder-Lowe Public-Key Protocol*}
 | 
| 
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 | theory Guard_NS_Public imports Guard_Public begin | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 11 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 12 | 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 | 13 | |
| 20768 | 14 | abbreviation (input) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 15 | ns1 :: "agent => agent => nat => event" where | 
| 20768 | 16 |   "ns1 A B NA == Says A B (Crypt (pubK B) {|Nonce NA, Agent A|})"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 17 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 18 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 19 | ns1' :: "agent => agent => agent => nat => event" where | 
| 20768 | 20 |   "ns1' A' A B NA == Says A' B (Crypt (pubK B) {|Nonce NA, Agent A|})"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 21 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 22 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 23 | ns2 :: "agent => agent => nat => nat => event" where | 
| 20768 | 24 |   "ns2 B A NA NB == Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 25 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 26 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 27 | ns2' :: "agent => agent => agent => nat => nat => event" where | 
| 20768 | 28 |   "ns2' B' B A NA NB == Says B' A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 29 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 30 | abbreviation (input) | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 31 | ns3 :: "agent => agent => nat => event" where | 
| 20768 | 32 | "ns3 A B NB == Says A B (Crypt (pubK B) (Nonce NB))" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 33 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 34 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 35 | 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 | 36 | |
| 23746 | 37 | inductive_set nsp :: "event list set" | 
| 38 | where | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 39 | |
| 23746 | 40 | Nil: "[]:nsp" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 41 | |
| 23746 | 42 | | Fake: "[| evs:nsp; X:synth (analz (spies evs)) |] ==> Says Spy B X # evs : nsp" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 43 | |
| 23746 | 44 | | NS1: "[| evs1:nsp; Nonce NA ~:used evs1 |] ==> ns1 A B NA # evs1 : nsp" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 45 | |
| 23746 | 46 | | NS2: "[| evs2:nsp; Nonce NB ~:used evs2; ns1' A' A B NA:set evs2 |] ==> | 
| 47 | ns2 B A NA NB # evs2:nsp" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 48 | |
| 23746 | 49 | | NS3: "!!A B B' NA NB evs3. [| evs3:nsp; ns1 A B NA:set evs3; ns2' B' B A NA NB:set evs3 |] ==> | 
| 50 | ns3 A B NB # evs3:nsp" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 51 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 52 | 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 | 53 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 54 | 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 | 55 | 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 | 56 | 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 | 57 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 58 | subsection{*general properties of nsp*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 59 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 60 | lemma nsp_has_no_Gets: "evs:nsp ==> 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 | 61 | by (erule nsp.induct, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 62 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 63 | lemma nsp_is_Gets_correct [iff]: "Gets_correct nsp" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 64 | by (auto simp: Gets_correct_def dest: nsp_has_no_Gets) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 65 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 66 | lemma nsp_is_one_step [iff]: "one_step nsp" | 
| 23746 | 67 | by (unfold one_step_def, clarify, ind_cases "ev#evs:nsp" 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 | 68 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 69 | lemma nsp_has_only_Says' [rule_format]: "evs:nsp ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 70 | 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 | 71 | by (erule nsp.induct, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 72 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 73 | lemma nsp_has_only_Says [iff]: "has_only_Says nsp" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 74 | by (auto simp: has_only_Says_def dest: nsp_has_only_Says') | 
| 
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 nsp_is_regular [iff]: "regular nsp" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 77 | 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 | 78 | by (erule nsp.induct, auto simp: 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 | 79 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 80 | subsection{*nonce are used only once*}
 | 
| 
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 NA_is_uniq [rule_format]: "evs:nsp ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 83 | Crypt (pubK B) {|Nonce NA, Agent A|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 84 | --> Crypt (pubK B') {|Nonce NA, Agent A'|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 85 | --> Nonce NA ~:analz (spies evs) --> A=A' & B=B'" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 86 | apply (erule nsp.induct, simp_all) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 87 | by (blast intro: analz_insertI)+ | 
| 
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 no_Nonce_NS1_NS2 [rule_format]: "evs:nsp ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 90 | Crypt (pubK B') {|Nonce NA', Nonce NA, Agent A'|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 91 | --> Crypt (pubK B) {|Nonce NA, Agent A|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 92 | --> Nonce NA:analz (spies evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 93 | apply (erule nsp.induct, simp_all) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 94 | by (blast intro: analz_insertI)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 95 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 96 | lemma no_Nonce_NS1_NS2' [rule_format]: | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 97 | "[| Crypt (pubK B') {|Nonce NA', Nonce NA, Agent A'|}:parts (spies evs);
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 98 | Crypt (pubK B) {|Nonce NA, Agent A|}:parts (spies evs); evs:nsp |]
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 99 | ==> Nonce NA:analz (spies evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 100 | by (rule no_Nonce_NS1_NS2, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 101 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 102 | lemma NB_is_uniq [rule_format]: "evs:nsp ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 103 | Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 104 | --> Crypt (pubK A') {|Nonce NA', Nonce NB, Agent B'|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 105 | --> Nonce NB ~:analz (spies evs) --> 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 | 106 | apply (erule nsp.induct, simp_all) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 107 | by (blast intro: analz_insertI)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 108 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 109 | subsection{*guardedness of NA*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 110 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 111 | lemma ns1_imp_Guard [rule_format]: "[| evs:nsp; 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 | 112 | ns1 A B NA:set evs --> Guard NA {priK A,priK B} (spies evs)"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 113 | apply (erule nsp.induct) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 114 | (* Nil *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 115 | apply simp_all | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 116 | (* Fake *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 117 | apply safe | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 118 | apply (erule in_synth_Guard, erule Guard_analz, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 119 | (* NS1 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 120 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 121 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 122 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 123 | apply (drule 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 | 124 | (* NS2 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 125 | apply (frule_tac A=A 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 | 126 | apply (case_tac "NAa=NA") | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 127 | apply (drule Guard_Nonce_analz, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 128 | apply (drule Says_imp_knows_Spy)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 129 | apply (drule_tac B=B and A'=Aa in NA_is_uniq, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 130 | (* NS3 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 131 | apply (case_tac "NB=NA", clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 132 | apply (drule Guard_Nonce_analz, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 133 | apply (drule Says_imp_knows_Spy)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 134 | by (drule no_Nonce_NS1_NS2, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 135 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 136 | 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 | 137 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 138 | lemma ns2_imp_Guard [rule_format]: "[| evs:nsp; 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 | 139 | ns2 B A NA NB:set evs --> Guard NB {priK A,priK B} (spies evs)" 
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 140 | apply (erule nsp.induct) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 141 | (* Nil *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 142 | apply simp_all | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 143 | (* Fake *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 144 | apply safe | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 145 | apply (erule in_synth_Guard, erule Guard_analz, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 146 | (* NS1 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 147 | apply (frule Nonce_neq, simp+, blast, 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 | 148 | (* NS2 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 149 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 150 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 151 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 152 | apply (frule_tac A=B and 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 | 153 | apply (case_tac "NAa=NB") | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 154 | apply (drule Guard_Nonce_analz, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 155 | apply (drule Says_imp_knows_Spy)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 156 | apply (drule no_Nonce_NS1_NS2, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 157 | (* NS3 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 158 | 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 | 159 | apply (drule Guard_Nonce_analz, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 160 | apply (drule Says_imp_knows_Spy)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 161 | by (drule_tac A=Aa and A'=A in NB_is_uniq, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 162 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 163 | subsection{*Agents' Authentication*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 164 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 165 | lemma B_trusts_NS1: "[| evs:nsp; 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 | 166 | Crypt (pubK B) {|Nonce NA, Agent A|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 167 | --> Nonce NA ~:analz (spies evs) --> ns1 A B NA:set evs" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 168 | apply (erule nsp.induct, simp_all) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 169 | by (blast intro: analz_insertI)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 170 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 171 | lemma A_trusts_NS2: "[| evs:nsp; A ~:bad; B ~:bad |] ==> ns1 A B NA:set evs | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 172 | --> Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}:parts (spies evs)
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 173 | --> ns2 B A 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 | 174 | apply (erule nsp.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 | 175 | apply (frule_tac B=B in ns1_imp_Guard, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 176 | apply (drule Guard_Nonce_analz, simp+, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 177 | apply (frule_tac B=B in ns1_imp_Guard, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 178 | apply (drule Guard_Nonce_analz, simp+, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 179 | apply (frule_tac B=B in ns1_imp_Guard, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 180 | by (drule Guard_Nonce_analz, simp+, blast+) | 
| 
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 B_trusts_NS3: "[| evs:nsp; A ~:bad; B ~:bad |] ==> ns2 B A 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 | 183 | --> Crypt (pubK B) (Nonce NB):parts (spies evs) --> ns3 A B NB:set evs" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 184 | apply (erule nsp.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 | 185 | apply (frule_tac B=B in ns2_imp_Guard, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 186 | apply (drule Guard_Nonce_analz, simp+, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 187 | apply (frule_tac B=B in ns2_imp_Guard, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 188 | apply (drule Guard_Nonce_analz, simp+, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 189 | apply (frule_tac B=B in ns2_imp_Guard, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 190 | apply (drule Guard_Nonce_analz, simp+, blast, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 191 | apply (frule_tac B=B in ns2_imp_Guard, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 192 | by (drule Guard_Nonce_analz, auto dest: Says_imp_knows_Spy NB_is_uniq) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 193 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 194 | end |