| author | paulson | 
| Fri, 03 Dec 2004 15:27:47 +0100 | |
| changeset 15369 | 090b16d6c6e0 | 
| parent 14945 | 7bfe4fa8a88f | 
| child 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 6452 | 1 | (* Title: HOL/Auth/KerberosIV | 
| 2 | ID: $Id$ | |
| 3 | Author: Giampaolo Bella, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 7 | header{*The Kerberos Protocol, Version IV*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 8 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 9 | theory KerberosIV = Public: | 
| 6452 | 10 | |
| 11 | syntax | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 12 | Kas :: agent | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 13 |   Tgs :: agent  --{*the two servers are translations...*}
 | 
| 6452 | 14 | |
| 15 | ||
| 16 | translations | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 17 | "Kas" == "Server " | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 18 | "Tgs" == "Friend 0" | 
| 6452 | 19 | |
| 20 | ||
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 21 | axioms | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 22 | Tgs_not_bad [iff]: "Tgs \<notin> bad" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 23 |    --{*Tgs is secure --- we already know that Kas is secure*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 24 | |
| 6452 | 25 | (*The current time is just the length of the trace!*) | 
| 26 | syntax | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 27 | CT :: "event list=>nat" | 
| 6452 | 28 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 29 | ExpirAuth :: "[nat, event list] => bool" | 
| 6452 | 30 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 31 | ExpirServ :: "[nat, event list] => bool" | 
| 6452 | 32 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 33 | ExpirAutc :: "[nat, event list] => bool" | 
| 6452 | 34 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 35 | RecentResp :: "[nat, nat] => bool" | 
| 6452 | 36 | |
| 37 | ||
| 38 | constdefs | |
| 39 | (* AuthKeys are those contained in an AuthTicket *) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 40 | AuthKeys :: "event list => key set" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 41 |     "AuthKeys evs == {AuthKey. \<exists>A Peer Tk. Says Kas A
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 42 |                         (Crypt (shrK A) {|Key AuthKey, Agent Peer, Tk,
 | 
| 6452 | 43 |                    (Crypt (shrK Peer) {|Agent A, Agent Peer, Key AuthKey, Tk|})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 44 | |}) \<in> set evs}" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 45 | |
| 6452 | 46 | (* A is the true creator of X if she has sent X and X never appeared on | 
| 47 | the trace before this event. Recall that traces grow from head. *) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 48 | Issues :: "[agent, agent, msg, event list] => bool" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 49 |              ("_ Issues _ with _ on _")
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 50 | "A Issues B with X on evs == | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 51 |       \<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 52 | X \<notin> parts (spies (takeWhile (% z. z \<noteq> Says A B Y) (rev evs)))" | 
| 6452 | 53 | |
| 54 | ||
| 55 | consts | |
| 56 | (*Duration of the authentication key*) | |
| 57 | AuthLife :: nat | |
| 58 | ||
| 59 | (*Duration of the service key*) | |
| 60 | ServLife :: nat | |
| 61 | ||
| 62 | (*Duration of an authenticator*) | |
| 63 | AutcLife :: nat | |
| 64 | ||
| 65 | (*Upper bound on the time of reaction of a server*) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 66 | RespLife :: nat | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 67 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 68 | specification (AuthLife) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 69 | AuthLife_LB [iff]: "2 \<le> AuthLife" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 70 | by blast | 
| 6452 | 71 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 72 | specification (ServLife) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 73 | ServLife_LB [iff]: "2 \<le> ServLife" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 74 | by blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 75 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 76 | specification (AutcLife) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 77 | AutcLife_LB [iff]: "Suc 0 \<le> AutcLife" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 78 | by blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 79 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 80 | specification (RespLife) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 81 | RespLife_LB [iff]: "Suc 0 \<le> RespLife" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 82 | by blast | 
| 6452 | 83 | |
| 84 | translations | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 85 | "CT" == "length " | 
| 6452 | 86 | |
| 87 | "ExpirAuth T evs" == "AuthLife + T < CT evs" | |
| 88 | ||
| 89 | "ExpirServ T evs" == "ServLife + T < CT evs" | |
| 90 | ||
| 91 | "ExpirAutc T evs" == "AutcLife + T < CT evs" | |
| 92 | ||
| 93 | "RecentResp T1 T2" == "T1 <= RespLife + T2" | |
| 94 | ||
| 95 | (*---------------------------------------------------------------------*) | |
| 96 | ||
| 97 | ||
| 98 | (* Predicate formalising the association between AuthKeys and ServKeys *) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 99 | constdefs | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 100 | KeyCryptKey :: "[key, key, event list] => bool" | 
| 6452 | 101 | "KeyCryptKey AuthKey ServKey evs == | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 102 | \<exists>A B tt. | 
| 6452 | 103 | Says Tgs A (Crypt AuthKey | 
| 104 |                      {|Key ServKey, Agent B, tt,
 | |
| 105 |                        Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} |})
 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 106 | \<in> set evs" | 
| 6452 | 107 | |
| 108 | consts | |
| 109 | ||
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 110 | kerberos :: "event list set" | 
| 6452 | 111 | inductive "kerberos" | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 112 | intros | 
| 6452 | 113 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 114 | Nil: "[] \<in> kerberos" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 115 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 116 | Fake: "[| evsf \<in> kerberos; X \<in> synth (analz (spies evsf)) |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 117 | ==> Says Spy B X # evsf \<in> kerberos" | 
| 6452 | 118 | |
| 119 | (* FROM the initiator *) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 120 | K1: "[| evs1 \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 121 |           ==> Says A Kas {|Agent A, Agent Tgs, Number (CT evs1)|} # evs1
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 122 | \<in> kerberos" | 
| 6452 | 123 | |
| 124 | (* Adding the timestamp serves to A in K3 to check that | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 125 | she doesn't get a reply too late. This kind of timeouts are ordinary. | 
| 6452 | 126 | If a server's reply is late, then it is likely to be fake. *) | 
| 127 | ||
| 128 | (*---------------------------------------------------------------------*) | |
| 129 | ||
| 130 | (*FROM Kas *) | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 131 | K2: "[| evs2 \<in> kerberos; Key AuthKey \<notin> used evs2; AuthKey \<in> symKeys; | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 132 |             Says A' Kas {|Agent A, Agent Tgs, Number Ta|} \<in> set evs2 |]
 | 
| 6452 | 133 | ==> Says Kas A | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 134 |                 (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number (CT evs2),
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 135 |                       (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 136 | Number (CT evs2)|})|}) # evs2 \<in> kerberos" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 137 | (* | 
| 6452 | 138 | The internal encryption builds the AuthTicket. | 
| 139 | The timestamp doesn't change inside the two encryptions: the external copy | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 140 | will be used by the initiator in K3; the one inside the | 
| 6452 | 141 | AuthTicket by Tgs in K4. | 
| 142 | *) | |
| 143 | ||
| 144 | (*---------------------------------------------------------------------*) | |
| 145 | ||
| 146 | (* FROM the initiator *) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 147 | K3: "[| evs3 \<in> kerberos; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 148 |             Says A Kas {|Agent A, Agent Tgs, Number Ta|} \<in> set evs3;
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 149 |             Says Kas' A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 150 | AuthTicket|}) \<in> set evs3; | 
| 6452 | 151 | RecentResp Tk Ta | 
| 152 | |] | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 153 |           ==> Says A Tgs {|AuthTicket,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 154 |                            (Crypt AuthKey {|Agent A, Number (CT evs3)|}),
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 155 | Agent B|} # evs3 \<in> kerberos" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 156 | (*The two events amongst the premises allow A to accept only those AuthKeys | 
| 6452 | 157 | that are not issued late. *) | 
| 158 | ||
| 159 | (*---------------------------------------------------------------------*) | |
| 160 | ||
| 161 | (* FROM Tgs *) | |
| 162 | (* Note that the last temporal check is not mentioned in the original MIT | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 163 | specification. Adding it strengthens the guarantees assessed by the | 
| 6452 | 164 | protocol. Theorems that exploit it have the suffix `_refined' | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 165 | *) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 166 | K4: "[| evs4 \<in> kerberos; Key ServKey \<notin> used evs4; ServKey \<in> symKeys; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 167 | B \<noteq> Tgs; AuthKey \<in> symKeys; | 
| 6452 | 168 |             Says A' Tgs {|
 | 
| 169 |              (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey,
 | |
| 170 | Number Tk|}), | |
| 171 |              (Crypt AuthKey {|Agent A, Number Ta1|}), Agent B|}
 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 172 | \<in> set evs4; | 
| 6452 | 173 | ~ ExpirAuth Tk evs4; | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 174 | ~ ExpirAutc Ta1 evs4; | 
| 6452 | 175 | ServLife + (CT evs4) <= AuthLife + Tk | 
| 176 | |] | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 177 | ==> Says Tgs A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 178 |                 (Crypt AuthKey {|Key ServKey, Agent B, Number (CT evs4),
 | 
| 6452 | 179 | 			       Crypt (shrK B) {|Agent A, Agent B, Key ServKey,
 | 
| 180 | Number (CT evs4)|} |}) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 181 | # evs4 \<in> kerberos" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 182 | (* Tgs creates a new session key per each request for a service, without | 
| 6452 | 183 | checking if there is still a fresh one for that service. | 
| 184 | The cipher under Tgs' key is the AuthTicket, the cipher under B's key | |
| 185 | is the ServTicket, which is built now. | |
| 186 | NOTE that the last temporal check is not present in the MIT specification. | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 187 | |
| 6452 | 188 | *) | 
| 189 | ||
| 190 | (*---------------------------------------------------------------------*) | |
| 191 | ||
| 192 | (* FROM the initiator *) | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 193 | K5: "[| evs5 \<in> kerberos; AuthKey \<in> symKeys; ServKey \<in> symKeys; | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 194 | Says A Tgs | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 195 |                 {|AuthTicket, Crypt AuthKey {|Agent A, Number Ta1|},
 | 
| 6452 | 196 | Agent B|} | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 197 | \<in> set evs5; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 198 | Says Tgs' A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 199 |              (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 200 | \<in> set evs5; | 
| 6452 | 201 | RecentResp Tt Ta1 |] | 
| 202 |           ==> Says A B {|ServTicket,
 | |
| 203 | 			 Crypt ServKey {|Agent A, Number (CT evs5)|} |}
 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 204 | # evs5 \<in> kerberos" | 
| 6452 | 205 | (* Checks similar to those in K3. *) | 
| 206 | ||
| 207 | (*---------------------------------------------------------------------*) | |
| 208 | ||
| 209 | (* FROM the responder*) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 210 | K6: "[| evs6 \<in> kerberos; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 211 |             Says A' B {|
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 212 |               (Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}),
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 213 |               (Crypt ServKey {|Agent A, Number Ta2|})|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 214 | \<in> set evs6; | 
| 6452 | 215 | ~ ExpirServ Tt evs6; | 
| 216 | ~ ExpirAutc Ta2 evs6 | |
| 217 | |] | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 218 | ==> Says B A (Crypt ServKey (Number Ta2)) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 219 | # evs6 \<in> kerberos" | 
| 6452 | 220 | (* Checks similar to those in K4. *) | 
| 221 | ||
| 222 | (*---------------------------------------------------------------------*) | |
| 223 | ||
| 224 | (* Leaking an AuthKey... *) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 225 | Oops1: "[| evsO1 \<in> kerberos; A \<noteq> Spy; | 
| 6452 | 226 | Says Kas A | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 227 |                 (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 228 | AuthTicket|}) \<in> set evsO1; | 
| 6452 | 229 | ExpirAuth Tk evsO1 |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 230 |           ==> Says A Spy {|Agent A, Agent Tgs, Number Tk, Key AuthKey|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 231 | # evsO1 \<in> kerberos" | 
| 6452 | 232 | |
| 233 | (*---------------------------------------------------------------------*) | |
| 234 | ||
| 235 | (*Leaking a ServKey... *) | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 236 | Oops2: "[| evsO2 \<in> kerberos; A \<noteq> Spy; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 237 | Says Tgs A | 
| 6452 | 238 |                 (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 239 | \<in> set evsO2; | 
| 6452 | 240 | ExpirServ Tt evsO2 |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 241 |           ==> Says A Spy {|Agent A, Agent B, Number Tt, Key ServKey|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 242 | # evsO2 \<in> kerberos" | 
| 6452 | 243 | |
| 244 | (*---------------------------------------------------------------------*) | |
| 245 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 246 | declare Says_imp_knows_Spy [THEN parts.Inj, dest] | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 247 | declare parts.Body [dest] | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 248 | declare analz_into_parts [dest] | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 249 | declare Fake_parts_insert_in_Un [dest] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 250 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 251 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 252 | subsection{*Lemmas about Lists*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 253 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 254 | lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 255 | apply (induct_tac "evs") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 256 | apply (induct_tac [2] "a", auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 257 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 258 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 259 | lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 260 | apply (induct_tac "evs") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 261 | apply (induct_tac [2] "a", auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 262 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 263 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 264 | lemma spies_Notes_rev: "spies (evs @ [Notes A X]) = | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 265 | (if A:bad then insert X (spies evs) else spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 266 | apply (induct_tac "evs") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 267 | apply (induct_tac [2] "a", auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 268 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 269 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 270 | lemma spies_evs_rev: "spies evs = spies (rev evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 271 | apply (induct_tac "evs") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 272 | apply (induct_tac [2] "a") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 273 | apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 274 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 275 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 276 | lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 277 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 278 | lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 279 | apply (induct_tac "evs") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 280 | apply (induct_tac [2] "a", auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 281 | txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 282 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 283 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 284 | lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 285 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 286 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 287 | subsection{*Lemmas about @{term AuthKeys}*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 288 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 289 | lemma AuthKeys_empty: "AuthKeys [] = {}"
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 290 | apply (unfold AuthKeys_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 291 | apply (simp (no_asm)) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 292 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 293 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 294 | lemma AuthKeys_not_insert: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 295 | "(\<forall>A Tk akey Peer. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 296 |    ev \<noteq> Says Kas A (Crypt (shrK A) {|akey, Agent Peer, Tk,
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 297 |               (Crypt (shrK Peer) {|Agent A, Agent Peer, akey, Tk|})|}))
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 298 | ==> AuthKeys (ev # evs) = AuthKeys evs" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 299 | by (unfold AuthKeys_def, auto) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 300 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 301 | lemma AuthKeys_insert: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 302 | "AuthKeys | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 303 |      (Says Kas A (Crypt (shrK A) {|Key K, Agent Peer, Number Tk,
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 304 |       (Crypt (shrK Peer) {|Agent A, Agent Peer, Key K, Number Tk|})|}) # evs)
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 305 | = insert K (AuthKeys evs)" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 306 | by (unfold AuthKeys_def, auto) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 307 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 308 | lemma AuthKeys_simp: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 309 | "K \<in> AuthKeys | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 310 |     (Says Kas A (Crypt (shrK A) {|Key K', Agent Peer, Number Tk,
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 311 |      (Crypt (shrK Peer) {|Agent A, Agent Peer, Key K', Number Tk|})|}) # evs)
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 312 | ==> K = K' | K \<in> AuthKeys evs" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 313 | by (unfold AuthKeys_def, auto) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 314 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 315 | lemma AuthKeysI: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 316 |    "Says Kas A (Crypt (shrK A) {|Key K, Agent Tgs, Number Tk,
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 317 |      (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key K, Number Tk|})|}) \<in> set evs
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 318 | ==> K \<in> AuthKeys evs" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 319 | by (unfold AuthKeys_def, auto) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 320 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 321 | lemma AuthKeys_used: "K \<in> AuthKeys evs ==> Key K \<in> used evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 322 | by (simp add: AuthKeys_def, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 323 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 324 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 325 | subsection{*Forwarding Lemmas*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 326 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 327 | text{*--For reasoning about the encrypted portion of message K3--*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 328 | lemma K3_msg_in_parts_spies: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 329 |      "Says Kas' A (Crypt KeyA {|AuthKey, Peer, Tk, AuthTicket|})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 330 | \<in> set evs ==> AuthTicket \<in> parts (spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 331 | by blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 332 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 333 | lemma Oops_range_spies1: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 334 |      "[| Says Kas A (Crypt KeyA {|Key AuthKey, Peer, Tk, AuthTicket|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 335 | \<in> set evs ; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 336 | evs \<in> kerberos |] ==> AuthKey \<notin> range shrK & AuthKey \<in> symKeys" | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 337 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 338 | apply (erule kerberos.induct, auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 339 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 340 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 341 | text{*--For reasoning about the encrypted portion of message K5--*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 342 | lemma K5_msg_in_parts_spies: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 343 |      "Says Tgs' A (Crypt AuthKey {|ServKey, Agent B, Tt, ServTicket|})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 344 | \<in> set evs ==> ServTicket \<in> parts (spies evs)" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 345 | by blast | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 346 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 347 | lemma Oops_range_spies2: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 348 |      "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 349 | \<in> set evs ; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 350 | evs \<in> kerberos |] ==> ServKey \<notin> range shrK & ServKey \<in> symKeys" | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 351 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 352 | apply (erule kerberos.induct, auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 353 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 354 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 355 | lemma Says_ticket_in_parts_spies: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 356 |      "Says S A (Crypt K {|SesKey, B, TimeStamp, Ticket|}) \<in> set evs
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 357 | ==> Ticket \<in> parts (spies evs)" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 358 | by blast | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 359 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 360 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 361 | (*Spy never sees another agent's shared key! (unless it's lost at start)*) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 362 | lemma Spy_see_shrK [simp]: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 363 | "evs \<in> kerberos ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 364 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 365 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 366 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 367 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 368 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 369 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 370 | lemma Spy_analz_shrK [simp]: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 371 | "evs \<in> kerberos ==> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 372 | by auto | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 373 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 374 | lemma Spy_see_shrK_D [dest!]: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 375 | "[| Key (shrK A) \<in> parts (spies evs); evs \<in> kerberos |] ==> A:bad" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 376 | by (blast dest: Spy_see_shrK) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 377 | lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 378 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 379 | text{*Nobody can have used non-existent keys!*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 380 | lemma new_keys_not_used [simp]: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 381 | "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> kerberos|] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 382 | ==> K \<notin> keysFor (parts (spies evs))" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 383 | apply (erule rev_mp) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 384 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 385 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 386 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 387 | txt{*Fake*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 388 | apply (force dest!: keysFor_parts_insert) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 389 | txt{*Others*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 390 | apply (force dest!: analz_shrK_Decrypt)+ | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 391 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 392 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 393 | (*Earlier, all protocol proofs declared this theorem. | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 394 | But few of them actually need it! (Another is Yahalom) *) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 395 | lemma new_keys_not_analzd: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 396 | "[|evs \<in> kerberos; K \<in> symKeys; Key K \<notin> used evs|] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 397 | ==> K \<notin> keysFor (analz (spies evs))" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 398 | by (blast dest: new_keys_not_used intro: keysFor_mono [THEN subsetD]) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 399 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 400 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 401 | subsection{*Regularity Lemmas*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 402 | text{*These concern the form of items passed in messages*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 403 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 404 | text{*Describes the form of AuthKey, AuthTicket, and K sent by Kas*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 405 | lemma Says_Kas_message_form: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 406 |      "[| Says Kas A (Crypt K {|Key AuthKey, Agent Peer, Tk, AuthTicket|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 407 | \<in> set evs; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 408 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 409 | ==> AuthKey \<notin> range shrK & AuthKey \<in> AuthKeys evs & AuthKey \<in> symKeys & | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 410 |   AuthTicket = (Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}) &
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 411 | K = shrK A & Peer = Tgs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 412 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 413 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 414 | apply (simp_all (no_asm) add: AuthKeys_def AuthKeys_insert) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 415 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 416 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 417 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 418 | (*This lemma is essential for proving Says_Tgs_message_form: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 419 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 420 | the session key AuthKey | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 421 | supplied by Kas in the authentication ticket | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 422 | cannot be a long-term key! | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 423 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 424 | Generalised to any session keys (both AuthKey and ServKey). | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 425 | *) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 426 | lemma SesKey_is_session_key: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 427 |      "[| Crypt (shrK Tgs_B) {|Agent A, Agent Tgs_B, Key SesKey, Number T|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 428 | \<in> parts (spies evs); Tgs_B \<notin> bad; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 429 | evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 430 | ==> SesKey \<notin> range shrK" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 431 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 432 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 433 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 434 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 435 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 436 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 437 | lemma A_trusts_AuthTicket: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 438 |      "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 439 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 440 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 441 |       ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk,
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 442 |                  Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}|})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 443 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 444 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 445 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 446 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 447 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 448 | txt{*Fake, K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 449 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 450 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 451 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 452 | lemma AuthTicket_crypt_AuthKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 453 |      "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 454 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 455 | evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 456 | ==> AuthKey \<in> AuthKeys evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 457 | apply (frule A_trusts_AuthTicket, assumption) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 458 | apply (simp (no_asm) add: AuthKeys_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 459 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 460 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 461 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 462 | text{*Describes the form of ServKey, ServTicket and AuthKey sent by Tgs*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 463 | lemma Says_Tgs_message_form: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 464 |      "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 465 | \<in> set evs; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 466 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 467 | ==> B \<noteq> Tgs & | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 468 | ServKey \<notin> range shrK & ServKey \<notin> AuthKeys evs & ServKey \<in> symKeys & | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 469 |        ServTicket = (Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}) &
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 470 | AuthKey \<notin> range shrK & AuthKey \<in> AuthKeys evs & AuthKey \<in> symKeys" | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 471 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 472 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 473 | apply (simp_all add: AuthKeys_insert AuthKeys_not_insert AuthKeys_empty AuthKeys_simp, blast, auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 474 | txt{*Three subcases of Message 4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 475 | apply (blast dest!: AuthKeys_used Says_Kas_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 476 | apply (blast dest!: SesKey_is_session_key) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 477 | apply (blast dest: AuthTicket_crypt_AuthKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 478 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 479 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 480 | text{*Authenticity of AuthKey for A:
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 481 | If a certain encrypted message appears then it originated with Kas*} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 482 | lemma A_trusts_AuthKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 483 |      "[| Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 484 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 485 | A \<notin> bad; evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 486 |       ==> Says Kas A (Crypt (shrK A) {|Key AuthKey, Peer, Tk, AuthTicket|})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 487 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 488 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 489 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 490 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 491 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 492 | txt{*Fake*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 493 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 494 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 495 | apply (blast dest!: A_trusts_AuthTicket [THEN Says_Kas_message_form]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 496 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 497 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 498 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 499 | text{*If a certain encrypted message appears then it originated with Tgs*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 500 | lemma A_trusts_K4: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 501 |      "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 502 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 503 | Key AuthKey \<notin> analz (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 504 | AuthKey \<notin> range shrK; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 505 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 506 |  ==> \<exists>A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 507 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 508 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 509 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 510 | apply (erule kerberos.induct, analz_mono_contra) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 511 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 512 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 513 | txt{*Fake*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 514 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 515 | txt{*K2*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 516 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 517 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 518 | apply auto | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 519 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 520 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 521 | lemma AuthTicket_form: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 522 |      "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 523 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 524 | A \<notin> bad; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 525 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 526 | ==> AuthKey \<notin> range shrK & AuthKey \<in> symKeys & | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 527 |         AuthTicket = Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}"
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 528 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 529 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 530 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 531 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 532 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 533 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 534 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 535 | text{* This form holds also over an AuthTicket, but is not needed below.*}
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 536 | lemma ServTicket_form: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 537 |      "[| Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 538 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 539 | Key AuthKey \<notin> analz (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 540 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 541 | ==> ServKey \<notin> range shrK & ServKey \<in> symKeys & | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 542 |     (\<exists>A. ServTicket = Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})"
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 543 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 544 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 545 | apply (erule kerberos.induct, analz_mono_contra) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 546 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 547 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 548 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 549 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 550 | text{* Essentially the same as @{text AuthTicket_form} *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 551 | lemma Says_kas_message_form: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 552 | "[| Says Kas' A (Crypt (shrK A) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 553 |               {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \<in> set evs;
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 554 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 555 | ==> AuthKey \<notin> range shrK & AuthKey \<in> symKeys & | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 556 | AuthTicket = | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 557 |                   Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Tk|}
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 558 | | AuthTicket \<in> analz (spies evs)" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 559 | by (blast dest: analz_shrK_Decrypt AuthTicket_form | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 560 | Says_imp_spies [THEN analz.Inj]) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 561 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 562 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 563 | lemma Says_tgs_message_form: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 564 |  "[| Says Tgs' A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 565 | \<in> set evs; AuthKey \<in> symKeys; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 566 | evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 567 | ==> ServKey \<notin> range shrK & | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 568 | (\<exists>A. ServTicket = | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 569 | 	      Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 570 | | ServTicket \<in> analz (spies evs)" | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 571 | apply (frule Says_imp_spies [THEN analz.Inj], auto) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 572 | apply (force dest!: ServTicket_form) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 573 | apply (frule analz_into_parts) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 574 | apply (frule ServTicket_form, auto) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 575 | done | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 576 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 577 | subsection{*Unicity Theorems*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 578 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 579 | text{* The session key, if secure, uniquely identifies the Ticket
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 580 | whether AuthTicket or ServTicket. As a matter of fact, one can read | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 581 | also Tgs in the place of B. *} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 582 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 583 | lemma unique_CryptKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 584 |      "[| Crypt (shrK B)  {|Agent A,  Agent B,  Key SesKey, T|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 585 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 586 |          Crypt (shrK B') {|Agent A', Agent B', Key SesKey, T'|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 587 | \<in> parts (spies evs); Key SesKey \<notin> analz (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 588 | evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 589 | ==> A=A' & B=B' & T=T'" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 590 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 591 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 592 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 593 | apply (erule kerberos.induct, analz_mono_contra) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 594 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 595 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 596 | txt{*Fake, K2, K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 597 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 598 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 599 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 600 | text{*An AuthKey is encrypted by one and only one Shared key.
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 601 | A ServKey is encrypted by one and only one AuthKey. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 602 | *} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 603 | lemma Key_unique_SesKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 604 |      "[| Crypt K  {|Key SesKey,  Agent B, T, Ticket|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 605 | \<in> parts (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 606 |          Crypt K' {|Key SesKey,  Agent B', T', Ticket'|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 607 | \<in> parts (spies evs); Key SesKey \<notin> analz (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 608 | evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 609 | ==> K=K' & B=B' & T=T' & Ticket=Ticket'" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 610 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 611 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 612 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 613 | apply (erule kerberos.induct, analz_mono_contra) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 614 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 615 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 616 | txt{*Fake, K2, K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 617 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 618 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 619 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 620 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 621 | (* | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 622 | At reception of any message mentioning A, Kas associates shrK A with | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 623 | a new AuthKey. Realistic, as the user gets a new AuthKey at each login. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 624 | Similarly, at reception of any message mentioning an AuthKey | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 625 | (a legitimate user could make several requests to Tgs - by K3), Tgs | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 626 | associates it with a new ServKey. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 627 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 628 | Therefore, a goal like | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 629 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 630 | "evs \<in> kerberos | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 631 | ==> Key Kc \<notin> analz (spies evs) --> | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 632 | (\<exists>K' B' T' Ticket'. \<forall>K B T Ticket. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 633 |             Crypt Kc {|Key K, Agent B, T, Ticket|}
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 634 | \<in> parts (spies evs) --> K=K' & B=B' & T=T' & Ticket=Ticket')" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 635 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 636 | would fail on the K2 and K4 cases. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 637 | *) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 638 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 639 | lemma unique_AuthKeys: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 640 | "[| Says Kas A | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 641 |               (Crypt Ka {|Key AuthKey, Agent Tgs, Tk, X|}) \<in> set evs;
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 642 | Says Kas A' | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 643 |               (Crypt Ka' {|Key AuthKey, Agent Tgs, Tk', X'|}) \<in> set evs;
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 644 | evs \<in> kerberos |] ==> A=A' & Ka=Ka' & Tk=Tk' & X=X'" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 645 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 646 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 647 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 648 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 649 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 650 | txt{*K2*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 651 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 652 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 653 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 654 | text{* ServKey uniquely identifies the message from Tgs *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 655 | lemma unique_ServKeys: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 656 | "[| Says Tgs A | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 657 |               (Crypt K {|Key ServKey, Agent B, Tt, X|}) \<in> set evs;
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 658 | Says Tgs A' | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 659 |               (Crypt K' {|Key ServKey, Agent B', Tt', X'|}) \<in> set evs;
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 660 | evs \<in> kerberos |] ==> A=A' & B=B' & K=K' & Tt=Tt' & X=X'" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 661 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 662 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 663 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 664 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 665 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 666 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 667 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 668 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 669 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 670 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 671 | subsection{*Lemmas About the Predicate @{term KeyCryptKey}*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 672 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 673 | lemma not_KeyCryptKey_Nil [iff]: "~ KeyCryptKey AuthKey ServKey []" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 674 | by (simp add: KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 675 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 676 | lemma KeyCryptKeyI: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 677 |  "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |}) \<in> set evs;
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 678 | evs \<in> kerberos |] ==> KeyCryptKey AuthKey ServKey evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 679 | apply (unfold KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 680 | apply (blast dest: Says_Tgs_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 681 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 682 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 683 | lemma KeyCryptKey_Says [simp]: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 684 | "KeyCryptKey AuthKey ServKey (Says S A X # evs) = | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 685 | (Tgs = S & | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 686 | (\<exists>B tt. X = Crypt AuthKey | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 687 |                 {|Key ServKey, Agent B, tt,
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 688 |                   Crypt (shrK B) {|Agent A, Agent B, Key ServKey, tt|} |})
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 689 | | KeyCryptKey AuthKey ServKey evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 690 | apply (unfold KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 691 | apply (simp (no_asm)) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 692 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 693 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 694 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 695 | (*A fresh AuthKey cannot be associated with any other | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 696 | (with respect to a given trace). *) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 697 | lemma Auth_fresh_not_KeyCryptKey: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 698 | "[| Key AuthKey \<notin> used evs; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 699 | ==> ~ KeyCryptKey AuthKey ServKey evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 700 | apply (unfold KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 701 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 702 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 703 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 704 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 705 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 706 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 707 | (*A fresh ServKey cannot be associated with any other | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 708 | (with respect to a given trace). *) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 709 | lemma Serv_fresh_not_KeyCryptKey: | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 710 | "Key ServKey \<notin> used evs ==> ~ KeyCryptKey AuthKey ServKey evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 711 | apply (unfold KeyCryptKey_def, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 712 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 713 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 714 | lemma AuthKey_not_KeyCryptKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 715 |      "[| Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, tk|}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 716 | \<in> parts (spies evs); evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 717 | ==> ~ KeyCryptKey K AuthKey evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 718 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 719 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 720 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 721 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 722 | txt{*Fake*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 723 | apply blast | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 724 | txt{*K2: by freshness*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 725 | apply (simp add: KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 726 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 727 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 728 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 729 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 730 | text{*A secure serverkey cannot have been used to encrypt others*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 731 | lemma ServKey_not_KeyCryptKey: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 732 |  "[| Crypt (shrK B) {|Agent A, Agent B, Key SK, tt|} \<in> parts (spies evs);
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 733 | Key SK \<notin> analz (spies evs); SK \<in> symKeys; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 734 | B \<noteq> Tgs; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 735 | ==> ~ KeyCryptKey SK K evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 736 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 737 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 738 | apply (erule kerberos.induct, analz_mono_contra) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 739 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 740 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 741 | txt{*K4 splits into distinct subcases*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 742 | apply auto | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 743 | txt{*ServKey can't have been enclosed in two certificates*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 744 | prefer 2 apply (blast dest: unique_CryptKey) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 745 | txt{*ServKey is fresh and so could not have been used, by
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 746 |    @{text new_keys_not_used}*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 747 | apply (force dest!: Crypt_imp_invKey_keysFor simp add: KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 748 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 749 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 750 | text{*Long term keys are not issued as ServKeys*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 751 | lemma shrK_not_KeyCryptKey: | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 752 | "evs \<in> kerberos ==> ~ KeyCryptKey K (shrK A) evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 753 | apply (unfold KeyCryptKey_def) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 754 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 755 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 756 | apply (frule_tac [5] K3_msg_in_parts_spies, auto) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 757 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 758 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 759 | text{*The Tgs message associates ServKey with AuthKey and therefore not with any
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 760 | other key AuthKey.*} | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 761 | lemma Says_Tgs_KeyCryptKey: | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 762 |      "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, tt, X |})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 763 | \<in> set evs; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 764 | AuthKey' \<noteq> AuthKey; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 765 | ==> ~ KeyCryptKey AuthKey' ServKey evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 766 | apply (unfold KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 767 | apply (blast dest: unique_ServKeys) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 768 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 769 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 770 | lemma KeyCryptKey_not_KeyCryptKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 771 | "[| KeyCryptKey AuthKey ServKey evs; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 772 | ==> ~ KeyCryptKey ServKey K evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 773 | apply (erule rev_mp) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 774 | apply (erule kerberos.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 775 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 776 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, safe) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 777 | txt{*K4 splits into subcases*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 778 | apply simp_all | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 779 | prefer 4 apply (blast dest!: AuthKey_not_KeyCryptKey) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 780 | txt{*ServKey is fresh and so could not have been used, by
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 781 |    @{text new_keys_not_used}*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 782 | prefer 2 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 783 | apply (force dest!: Crypt_imp_invKey_keysFor simp add: KeyCryptKey_def) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 784 | txt{*Others by freshness*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 785 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 786 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 787 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 788 | text{*The only session keys that can be found with the help of session keys are
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 789 | those sent by Tgs in step K4. *} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 790 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 791 | text{*We take some pains to express the property
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 792 | as a logical equivalence so that the simplifier can apply it.*} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 793 | lemma Key_analz_image_Key_lemma: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 794 | "P --> (Key K \<in> analz (Key`KK Un H)) --> (K:KK | Key K \<in> analz H) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 795 | ==> | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 796 | P --> (Key K \<in> analz (Key`KK Un H)) = (K:KK | Key K \<in> analz H)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 797 | by (blast intro: analz_mono [THEN subsetD]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 798 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 799 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 800 | lemma KeyCryptKey_analz_insert: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 801 | "[| KeyCryptKey K K' evs; K \<in> symKeys; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 802 | ==> Key K' \<in> analz (insert (Key K) (spies evs))" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 803 | apply (simp add: KeyCryptKey_def, clarify) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 804 | apply (drule Says_imp_spies [THEN analz.Inj, THEN analz_insertI], auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 805 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 806 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 807 | lemma AuthKeys_are_not_KeyCryptKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 808 | "[| K \<in> AuthKeys evs Un range shrK; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 809 | ==> \<forall>SK. ~ KeyCryptKey SK K evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 810 | apply (simp add: KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 811 | apply (blast dest: Says_Tgs_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 812 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 813 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 814 | lemma not_AuthKeys_not_KeyCryptKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 815 | "[| K \<notin> AuthKeys evs; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 816 | K \<notin> range shrK; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 817 | ==> \<forall>SK. ~ KeyCryptKey K SK evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 818 | apply (simp add: KeyCryptKey_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 819 | apply (blast dest: Says_Tgs_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 820 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 821 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 822 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 823 | subsection{*Secrecy Theorems*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 824 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 825 | text{*For the Oops2 case of the next theorem*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 826 | lemma Oops2_not_KeyCryptKey: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 827 | "[| evs \<in> kerberos; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 828 | Says Tgs A (Crypt AuthKey | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 829 |                      {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 830 | \<in> set evs |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 831 | ==> ~ KeyCryptKey ServKey SK evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 832 | apply (blast dest: KeyCryptKeyI KeyCryptKey_not_KeyCryptKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 833 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 834 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 835 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 836 | text{* Big simplification law for keys SK that are not crypted by keys in KK
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 837 | It helps prove three, otherwise hard, facts about keys. These facts are | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 838 | exploited as simplification laws for analz, and also "limit the damage" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 839 | in case of loss of a key to the spy. See ESORICS98. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 840 | [simplified by LCP] *} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 841 | lemma Key_analz_image_Key [rule_format (no_asm)]: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 842 | "evs \<in> kerberos ==> | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 843 | (\<forall>SK KK. SK \<in> symKeys & KK <= -(range shrK) --> | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 844 | (\<forall>K \<in> KK. ~ KeyCryptKey K SK evs) --> | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 845 | (Key SK \<in> analz (Key`KK Un (spies evs))) = | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 846 | (SK \<in> KK | Key SK \<in> analz (spies evs)))" | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 847 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 848 | apply (frule_tac [10] Oops_range_spies2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 849 | apply (frule_tac [9] Oops_range_spies1) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 850 | apply (frule_tac [7] Says_tgs_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 851 | apply (frule_tac [5] Says_kas_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 852 | apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 853 | txt{*Case-splits for Oops1 and message 5: the negated case simplifies using
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 854 | the induction hypothesis*} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 855 | apply (case_tac [11] "KeyCryptKey AuthKey SK evsO1") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 856 | apply (case_tac [8] "KeyCryptKey ServKey SK evs5") | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 857 | apply (simp_all del: image_insert | 
| 14945 | 858 | add: analz_image_freshK_simps KeyCryptKey_Says) | 
| 859 | txt{*Fake*} 
 | |
| 860 | apply spy_analz | |
| 861 | apply (simp_all del: image_insert | |
| 862 | add: shrK_not_KeyCryptKey | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 863 | Oops2_not_KeyCryptKey Auth_fresh_not_KeyCryptKey | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 864 | Serv_fresh_not_KeyCryptKey Says_Tgs_KeyCryptKey Spy_analz_shrK) | 
| 14945 | 865 |   --{*Splitting the @{text simp_all} into two parts makes it faster.*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 866 | txt{*K2*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 867 | apply blast | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 868 | txt{*K3*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 869 | apply blast | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 870 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 871 | apply (blast dest!: AuthKey_not_KeyCryptKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 872 | txt{*K5*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 873 | apply (case_tac "Key ServKey \<in> analz (spies evs5) ") | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 874 | txt{*If ServKey is compromised then the result follows directly...*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 875 | apply (simp (no_asm_simp) add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 876 | txt{*...therefore ServKey is uncompromised.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 877 | txt{*The KeyCryptKey ServKey SK evs5 case leads to a contradiction.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 878 | apply (blast elim!: ServKey_not_KeyCryptKey [THEN [2] rev_notE] del: allE ballE) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 879 | txt{*Another K5 case*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 880 | apply blast | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 881 | txt{*Oops1*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 882 | apply simp | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 883 | apply (blast dest!: KeyCryptKey_analz_insert) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 884 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 885 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 886 | text{* First simplification law for analz: no session keys encrypt
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 887 | authentication keys or shared keys. *} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 888 | lemma analz_insert_freshK1: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 889 | "[| evs \<in> kerberos; K \<in> (AuthKeys evs) Un range shrK; | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 890 | K \<in> symKeys; | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 891 | SesKey \<notin> range shrK |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 892 | ==> (Key K \<in> analz (insert (Key SesKey) (spies evs))) = | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 893 | (K = SesKey | Key K \<in> analz (spies evs))" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 894 | apply (frule AuthKeys_are_not_KeyCryptKey, assumption) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 895 | apply (simp del: image_insert | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 896 | add: analz_image_freshK_simps add: Key_analz_image_Key) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 897 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 898 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 899 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 900 | text{* Second simplification law for analz: no service keys encrypt any other keys.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 901 | lemma analz_insert_freshK2: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 902 | "[| evs \<in> kerberos; ServKey \<notin> (AuthKeys evs); ServKey \<notin> range shrK; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 903 | K \<in> symKeys|] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 904 | ==> (Key K \<in> analz (insert (Key ServKey) (spies evs))) = | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 905 | (K = ServKey | Key K \<in> analz (spies evs))" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 906 | apply (frule not_AuthKeys_not_KeyCryptKey, assumption, assumption) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 907 | apply (simp del: image_insert | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 908 | add: analz_image_freshK_simps add: Key_analz_image_Key) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 909 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 910 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 911 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 912 | text{* Third simplification law for analz: only one authentication key encrypts a
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 913 | certain service key.*} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 914 | lemma analz_insert_freshK3: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 915 | "[| Says Tgs A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 916 |             (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 917 | \<in> set evs; ServKey \<in> symKeys; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 918 | AuthKey \<noteq> AuthKey'; AuthKey' \<notin> range shrK; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 919 | ==> (Key ServKey \<in> analz (insert (Key AuthKey') (spies evs))) = | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 920 | (ServKey = AuthKey' | Key ServKey \<in> analz (spies evs))" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 921 | apply (drule_tac AuthKey' = AuthKey' in Says_Tgs_KeyCryptKey, blast, assumption) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 922 | apply (simp del: image_insert | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 923 | add: analz_image_freshK_simps add: Key_analz_image_Key) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 924 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 925 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 926 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 927 | text{*a weakness of the protocol*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 928 | lemma AuthKey_compromises_ServKey: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 929 | "[| Says Tgs A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 930 |               (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 931 | \<in> set evs; AuthKey \<in> symKeys; | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 932 | Key AuthKey \<in> analz (spies evs); evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 933 | ==> Key ServKey \<in> analz (spies evs)" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 934 | by (force dest: Says_imp_spies [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Fst]) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 935 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 936 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 937 | subsection{* Guarantees for Kas *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 938 | lemma ServKey_notin_AuthKeysD: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 939 |      "[| Crypt AuthKey {|Key ServKey, Agent B, Tt,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 940 |                       Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 941 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 942 | Key ServKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 943 | B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 944 | ==> ServKey \<notin> AuthKeys evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 945 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 946 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 947 | apply (simp add: AuthKeys_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 948 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 949 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 950 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 951 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 952 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 953 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 954 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 955 | text{*If Spy sees the Authentication Key sent in msg K2, then
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 956 | the Key has expired.*} | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 957 | lemma Confidentiality_Kas_lemma [rule_format]: | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 958 | "[| AuthKey \<in> symKeys; A \<notin> bad; evs \<in> kerberos |] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 959 | ==> Says Kas A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 960 | (Crypt (shrK A) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 961 |                   {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 962 |           Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 963 | \<in> set evs --> | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 964 | Key AuthKey \<in> analz (spies evs) --> | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 965 | ExpirAuth Tk evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 966 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 967 | apply (frule_tac [10] Oops_range_spies2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 968 | apply (frule_tac [9] Oops_range_spies1) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 969 | apply (frule_tac [7] Says_tgs_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 970 | apply (frule_tac [5] Says_kas_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 971 | apply (safe del: impI conjI impCE) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 972 | apply (simp_all (no_asm_simp) add: Says_Kas_message_form less_SucI analz_insert_eq not_parts_not_analz analz_insert_freshK1 pushes) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 973 | txt{*Fake*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 974 | apply spy_analz | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 975 | txt{*K2*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 976 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 977 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 978 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 979 | txt{*Level 8: K5*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 980 | apply (blast dest: ServKey_notin_AuthKeysD Says_Kas_message_form intro: less_SucI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 981 | txt{*Oops1*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 982 | apply (blast dest!: unique_AuthKeys intro: less_SucI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 983 | txt{*Oops2*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 984 | apply (blast dest: Says_Tgs_message_form Says_Kas_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 985 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 986 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 987 | lemma Confidentiality_Kas: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 988 | "[| Says Kas A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 989 |               (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 990 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 991 | ~ ExpirAuth Tk evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 992 | A \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 993 | ==> Key AuthKey \<notin> analz (spies evs)" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 994 | by (blast dest: Says_Kas_message_form Confidentiality_Kas_lemma) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 995 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 996 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 997 | subsection{* Guarantees for Tgs *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 998 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 999 | text{*If Spy sees the Service Key sent in msg K4, then
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1000 | the Key has expired.*} | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1001 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1002 | lemma Confidentiality_lemma [rule_format]: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1003 | "[| Says Tgs A | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1004 | (Crypt AuthKey | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1005 | 	       {|Key ServKey, Agent B, Number Tt,
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1006 | 		 Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1007 | \<in> set evs; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1008 | Key AuthKey \<notin> analz (spies evs); | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1009 | ServKey \<in> symKeys; | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1010 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1011 | ==> Key ServKey \<in> analz (spies evs) --> | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1012 | ExpirServ Tt evs" | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1013 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1014 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1015 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1016 | apply (rule_tac [9] impI)+; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1017 |   --{*The Oops1 case is unusual: must simplify
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1018 |     @{term "Authkey \<notin> analz (spies (ev#evs))"}, not letting
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1019 |    @{text analz_mono_contra} weaken it to
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1020 |    @{term "Authkey \<notin> analz (spies evs)"},
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1021 |   for we then conclude @{term "AuthKey \<noteq> AuthKeya"}.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1022 | apply analz_mono_contra | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1023 | apply (frule_tac [10] Oops_range_spies2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1024 | apply (frule_tac [9] Oops_range_spies1) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1025 | apply (frule_tac [7] Says_tgs_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1026 | apply (frule_tac [5] Says_kas_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1027 | apply (safe del: impI conjI impCE) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1028 | apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3 pushes) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1029 | txt{*Fake*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1030 | apply spy_analz | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1031 | txt{*K2*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1032 | apply (blast intro: parts_insertI less_SucI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1033 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1034 | apply (blast dest: A_trusts_AuthTicket Confidentiality_Kas) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1035 | txt{*Oops2*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1036 | prefer 3 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1037 | apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1038 | txt{*Oops1*}
 | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1039 | prefer 2 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1040 | apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1041 | txt{*K5.  Not clear how this step could be integrated with the main
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1042 | simplification step.*} | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1043 | apply clarify | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1044 | apply (erule_tac V = "Says Aa Tgs ?X \<in> set ?evs" in thin_rl) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1045 | apply (frule Says_imp_spies [THEN parts.Inj, THEN ServKey_notin_AuthKeysD]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1046 | apply (assumption, blast, assumption) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1047 | apply (simp add: analz_insert_freshK2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1048 | apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1049 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1050 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1051 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1052 | text{* In the real world Tgs can't check wheter AuthKey is secure! *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1053 | lemma Confidentiality_Tgs1: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1054 | "[| Says Tgs A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1055 |               (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1056 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1057 | Key AuthKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1058 | ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1059 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1060 | ==> Key ServKey \<notin> analz (spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1061 | apply (blast dest: Says_Tgs_message_form Confidentiality_lemma) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1062 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1063 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1064 | text{* In the real world Tgs CAN check what Kas sends! *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1065 | lemma Confidentiality_Tgs2: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1066 | "[| Says Kas A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1067 |                (Crypt Ka {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1068 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1069 | Says Tgs A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1070 |               (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1071 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1072 | ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1073 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1074 | ==> Key ServKey \<notin> analz (spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1075 | apply (blast dest!: Confidentiality_Kas Confidentiality_Tgs1) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1076 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1077 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1078 | text{*Most general form*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1079 | lemmas Confidentiality_Tgs3 = A_trusts_AuthTicket [THEN Confidentiality_Tgs2] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1080 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1081 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1082 | subsection{* Guarantees for Alice *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1083 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1084 | lemmas Confidentiality_Auth_A = A_trusts_AuthKey [THEN Confidentiality_Kas] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1085 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1086 | lemma A_trusts_K4_bis: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1087 | "[| Says Kas A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1088 |        (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Tk, AuthTicket|}) \<in> set evs;
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1089 |      Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1090 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1091 | Key AuthKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1092 | evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1093 |  ==> Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1094 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1095 | apply (frule Says_Kas_message_form, assumption) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1096 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1097 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1098 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1099 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1100 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1101 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1102 | txt{*K2 and K4 remain*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1103 | prefer 2 apply (blast dest!: unique_CryptKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1104 | apply (blast dest!: A_trusts_K4 Says_Tgs_message_form AuthKeys_used) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1105 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1106 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1107 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1108 | lemma Confidentiality_Serv_A: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1109 |      "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1110 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1111 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1112 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1113 | ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1114 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1115 | ==> Key ServKey \<notin> analz (spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1116 | apply (drule A_trusts_AuthKey, assumption, assumption) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1117 | apply (blast dest: Confidentiality_Kas Says_Kas_message_form A_trusts_K4_bis Confidentiality_Tgs2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1118 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1119 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1120 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1121 | subsection{* Guarantees for Bob *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1122 | text{* Theorems for the refined model have suffix "refined"                *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1123 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1124 | lemma K4_imp_K2: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1125 | "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1126 | \<in> set evs; evs \<in> kerberos|] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1127 | ==> \<exists>Tk. Says Kas A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1128 | (Crypt (shrK A) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1129 |          {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1130 |            Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1131 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1132 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1133 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1134 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1135 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1136 | apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN A_trusts_AuthTicket]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1137 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1138 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1139 | lemma K4_imp_K2_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1140 | "[| Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1141 | \<in> set evs; evs \<in> kerberos|] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1142 |    ==> \<exists>Tk. (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1143 |            Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1144 | \<in> set evs | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1145 | & ServLife + Tt <= AuthLife + Tk)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1146 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1147 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1148 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1149 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, auto) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1150 | apply (blast dest!: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN A_trusts_AuthTicket]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1151 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1152 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1153 | text{*Authenticity of ServKey for B*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1154 | lemma B_trusts_ServKey: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1155 |      "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1156 | \<in> parts (spies evs); B \<noteq> Tgs; B \<notin> bad; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1157 | evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1158 | ==> \<exists>AuthKey. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1159 |        Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Tt,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1160 |                    Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Tt|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1161 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1162 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1163 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1164 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1165 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1166 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1167 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1168 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1169 | lemma B_trusts_ServTicket_Kas: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1170 |      "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1171 | \<in> parts (spies evs); B \<noteq> Tgs; B \<notin> bad; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1172 | evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1173 | ==> \<exists>AuthKey Tk. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1174 | Says Kas A | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1175 |          (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1176 |             Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1177 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1178 | by (blast dest!: B_trusts_ServKey K4_imp_K2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1179 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1180 | lemma B_trusts_ServTicket_Kas_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1181 |      "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1182 | \<in> parts (spies evs); B \<noteq> Tgs; B \<notin> bad; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1183 | evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1184 |   ==> \<exists>AuthKey Tk. Says Kas A (Crypt(shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1185 |            Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1186 | \<in> set evs | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1187 | & ServLife + Tt <= AuthLife + Tk" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1188 | by (blast dest!: B_trusts_ServKey K4_imp_K2_refined) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1189 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1190 | lemma B_trusts_ServTicket: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1191 |      "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1192 | \<in> parts (spies evs); B \<noteq> Tgs; B \<notin> bad; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1193 | evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1194 | ==> \<exists>Tk AuthKey. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1195 |      Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1196 |                    Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1197 | \<in> set evs | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1198 |      & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1199 |                    Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1200 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1201 | by (blast dest: B_trusts_ServKey K4_imp_K2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1202 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1203 | lemma B_trusts_ServTicket_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1204 |      "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1205 | \<in> parts (spies evs); B \<noteq> Tgs; B \<notin> bad; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1206 | evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1207 | ==> \<exists>Tk AuthKey. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1208 |      (Says Kas A (Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1209 |                    Crypt (shrK Tgs) {|Agent A, Agent Tgs, Key AuthKey, Number Tk|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1210 | \<in> set evs | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1211 |      & Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1212 |                    Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1213 | \<in> set evs | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1214 | & ServLife + Tt <= AuthLife + Tk)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1215 | by (blast dest: B_trusts_ServKey K4_imp_K2_refined) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1216 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1217 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1218 | lemma NotExpirServ_NotExpirAuth_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1219 | "[| ~ ExpirServ Tt evs; ServLife + Tt <= AuthLife + Tk |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1220 | ==> ~ ExpirAuth Tk evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1221 | by (blast dest: leI le_trans elim: leE) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1222 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1223 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1224 | lemma Confidentiality_B: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1225 |      "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1226 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1227 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1228 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1229 |          Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1230 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1231 | ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1232 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1233 | ==> Key ServKey \<notin> analz (spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1234 | apply (frule A_trusts_AuthKey) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1235 | apply (frule_tac [3] Confidentiality_Kas) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1236 | apply (frule_tac [6] B_trusts_ServTicket, auto) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1237 | apply (blast dest!: Confidentiality_Tgs2 dest: Says_Kas_message_form A_trusts_K4 unique_ServKeys unique_AuthKeys) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1238 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1239 | (* | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1240 | The proof above is fast. It can be done in one command in 17 secs: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1241 | apply (blast dest: A_trusts_AuthKey A_trusts_K4 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1242 | Says_Kas_message_form B_trusts_ServTicket | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1243 | unique_ServKeys unique_AuthKeys | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1244 | Confidentiality_Kas | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1245 | Confidentiality_Tgs2) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1246 | It is very brittle: we can't use this command partway | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1247 | through the script above. | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1248 | *) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1249 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1250 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1251 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1252 | text{*Most general form -- only for refined model! *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1253 | lemma Confidentiality_B_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1254 |      "[| Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1255 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1256 | ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1257 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1258 | ==> Key ServKey \<notin> analz (spies evs)" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1259 | apply (blast dest: B_trusts_ServTicket_refined NotExpirServ_NotExpirAuth_refined Confidentiality_Tgs2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1260 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1261 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1262 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1263 | subsection{* Authenticity theorems *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1264 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1265 | text{*1. Session Keys authenticity: they originated with servers.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1266 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1267 | text{*Authenticity of ServKey for A*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1268 | lemma A_trusts_ServKey: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1269 |      "[| Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1270 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1271 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1272 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1273 | ~ ExpirAuth Tk evs; A \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1274 |  ==>Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1275 | \<in> set evs" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1276 | by (blast dest: A_trusts_AuthKey Confidentiality_Auth_A A_trusts_K4_bis) | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1277 | |
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1278 | text{*Note: requires a temporal check*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1279 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1280 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1281 | (***2. Parties authenticity: each party verifies "the identity of | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1282 | another party who generated some data" (quoted from Neuman & Ts'o).***) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1283 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1284 | (*These guarantees don't assess whether two parties agree on | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1285 | the same session key: sending a message containing a key | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1286 | doesn't a priori state knowledge of the key.***) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1287 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1288 | text{*B checks authenticity of A by theorems @{text A_Authenticity} and
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1289 |   @{text A_authenticity_refined} *}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1290 | lemma Says_Auth: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1291 |      "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1292 |          Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1293 | ServTicket|}) \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1294 | Key ServKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1295 | A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1296 |  ==> Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|} \<in> set evs"
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1297 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1298 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1299 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1300 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1301 | apply (frule_tac [5] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1302 | apply (frule_tac [7] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1303 | apply (simp_all (no_asm_simp) add: all_conj_distrib) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1304 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1305 | txt{*K3*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1306 | apply (blast dest: A_trusts_AuthKey Says_Kas_message_form Says_Tgs_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1307 | txt{*K4*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1308 | apply (force dest!: Crypt_imp_keysFor) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1309 | txt{*K5*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1310 | apply (blast dest: Key_unique_SesKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1311 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1312 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1313 | text{*The second assumption tells B what kind of key ServKey is.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1314 | lemma A_Authenticity: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1315 |      "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1316 |          Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1317 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1318 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1319 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1320 |          Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1321 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1322 | ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1323 | B \<noteq> Tgs; A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1324 |    ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1325 |                   Crypt ServKey {|Agent A, Number Ta|} |} \<in> set evs"
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1326 | by (blast intro: Says_Auth dest: Confidentiality_B Key_unique_SesKey B_trusts_ServKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1327 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1328 | text{*Stronger form in the refined model*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1329 | lemma A_Authenticity_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1330 |      "[| Crypt ServKey {|Agent A, Number Ta2|} \<in> parts (spies evs);
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1331 |          Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1332 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1333 | ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1334 | B \<noteq> Tgs; A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1335 |    ==> Says A B {|Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|},
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1336 |                   Crypt ServKey {|Agent A, Number Ta2|} |} \<in> set evs"
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1337 | by (blast dest: Confidentiality_B_refined B_trusts_ServKey Key_unique_SesKey intro: Says_Auth) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1338 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1339 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1340 | text{*A checks authenticity of B by theorem @{text B_authenticity}*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1341 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1342 | lemma Says_K6: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1343 | "[| Crypt ServKey (Number Ta) \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1344 |          Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, Number Tt,
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1345 | ServTicket|}) \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1346 | Key ServKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1347 | A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1348 | ==> Says B A (Crypt ServKey (Number Ta)) \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1349 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1350 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1351 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1352 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1353 | apply (frule_tac [5] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1354 | apply (frule_tac [7] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1355 | apply (simp_all (no_asm_simp)) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1356 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1357 | apply (force dest!: Crypt_imp_keysFor, clarify) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1358 | apply (frule Says_Tgs_message_form, assumption, clarify) (*PROOF FAILED if omitted*) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1359 | apply (blast dest: unique_CryptKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1360 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1361 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1362 | lemma K4_trustworthy: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1363 |      "[| Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1364 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1365 | Key AuthKey \<notin> analz (spies evs); AuthKey \<notin> range shrK; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1366 | evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1367 |   ==> \<exists>A. Says Tgs A (Crypt AuthKey {|Key ServKey, Agent B, T, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1368 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1369 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1370 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1371 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1372 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1373 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1374 | apply (blast+) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1375 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1376 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1377 | lemma B_Authenticity: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1378 | "[| Crypt ServKey (Number Ta) \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1379 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1380 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1381 |          Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1382 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1383 | ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1384 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1385 | ==> Says B A (Crypt ServKey (Number Ta)) \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1386 | apply (frule A_trusts_AuthKey) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1387 | apply (frule_tac [3] Says_Kas_message_form) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1388 | apply (frule_tac [4] Confidentiality_Kas) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1389 | apply (frule_tac [7] K4_trustworthy) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1390 | prefer 8 apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1391 | apply (erule_tac [9] exE) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1392 | apply (frule_tac [9] K4_imp_K2) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1393 | txt{*Yes the proof's a mess, but I don't know how to improve it.*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1394 | apply assumption+ | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1395 | apply (blast dest: Key_unique_SesKey intro!: Says_K6 dest: Confidentiality_Tgs1 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1396 | ) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1397 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1398 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1399 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1400 | (***3. Parties' knowledge of session keys. A knows a session key if she | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1401 | used it to build a cipher.***) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1402 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1403 | lemma B_Knows_B_Knows_ServKey_lemma: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1404 | "[| Says B A (Crypt ServKey (Number Ta)) \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1405 | Key ServKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1406 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1407 | ==> B Issues A with (Crypt ServKey (Number Ta)) on evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1408 | apply (simp (no_asm) add: Issues_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1409 | apply (rule exI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1410 | apply (rule conjI, assumption) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1411 | apply (simp (no_asm)) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1412 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1413 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1414 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1415 | apply (frule_tac [5] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1416 | apply (frule_tac [7] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1417 | apply (simp_all (no_asm_simp) add: all_conj_distrib) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1418 | apply blast | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1419 | txt{*K6 requires numerous lemmas*}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1420 | apply (simp add: takeWhile_tail) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1421 | apply (blast dest: B_trusts_ServTicket parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: Says_K6) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1422 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1423 | (*Key ServKey \<notin> analz (spies evs) could be relaxed by Confidentiality_B | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1424 | but this is irrelevant because B knows what he knows! *) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1425 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1426 | lemma B_Knows_B_Knows_ServKey: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1427 | "[| Says B A (Crypt ServKey (Number Ta)) \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1428 |          Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1429 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1430 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1431 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1432 |          Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1433 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1434 | ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1435 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1436 | ==> B Issues A with (Crypt ServKey (Number Ta)) on evs" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1437 | by (blast dest!: Confidentiality_B B_Knows_B_Knows_ServKey_lemma) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1438 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1439 | lemma B_Knows_B_Knows_ServKey_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1440 | "[| Says B A (Crypt ServKey (Number Ta)) \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1441 |          Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1442 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1443 | ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1444 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1445 | ==> B Issues A with (Crypt ServKey (Number Ta)) on evs" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1446 | by (blast dest!: Confidentiality_B_refined B_Knows_B_Knows_ServKey_lemma) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1447 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1448 | lemma A_Knows_B_Knows_ServKey: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1449 | "[| Crypt ServKey (Number Ta) \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1450 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1451 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1452 |          Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1453 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1454 | ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1455 | A \<notin> bad; B \<notin> bad; B \<noteq> Tgs; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1456 | ==> B Issues A with (Crypt ServKey (Number Ta)) on evs" | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1457 | by (blast dest!: B_Authenticity Confidentiality_Serv_A B_Knows_B_Knows_ServKey_lemma) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1458 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1459 | lemma K3_imp_K2: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1460 | "[| Says A Tgs | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1461 |              {|AuthTicket, Crypt AuthKey {|Agent A, Number Ta|}, Agent B|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1462 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1463 | A \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1464 | ==> \<exists>Tk. Says Kas A (Crypt (shrK A) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1465 |                       {|Key AuthKey, Agent Tgs, Tk, AuthTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1466 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1467 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1468 | apply (erule kerberos.induct) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1469 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1470 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast, blast) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1471 | apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN A_trusts_AuthKey]) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1472 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1473 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1474 | lemma K4_trustworthy': | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1475 |      "[| Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1476 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1477 | Says Kas A (Crypt (shrK A) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1478 |                      {|Key AuthKey, Agent Tgs, Tk, AuthTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1479 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1480 | Key AuthKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1481 | B \<noteq> Tgs; A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1482 | ==> Says Tgs A (Crypt AuthKey | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1483 |                      {|Key ServKey, Agent B, Number Tt, ServTicket|})
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1484 | \<in> set evs" | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1485 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1486 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1487 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1488 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1489 | apply (frule_tac [7] K5_msg_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1490 | apply (frule_tac [5] K3_msg_in_parts_spies, simp_all, blast) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1491 | apply (force dest!: Crypt_imp_keysFor) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1492 | apply (blast dest: Says_imp_spies [THEN parts.Inj, THEN parts.Fst, THEN A_trusts_AuthTicket] unique_AuthKeys) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1493 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1494 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1495 | lemma A_Knows_A_Knows_ServKey_lemma: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1496 |      "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1497 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1498 | Key ServKey \<notin> analz (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1499 | B \<noteq> Tgs; A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1500 |    ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1501 | apply (simp (no_asm) add: Issues_def) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1502 | apply (rule exI) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1503 | apply (rule conjI, assumption) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1504 | apply (simp (no_asm)) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1505 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1506 | apply (erule rev_mp) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1507 | apply (erule kerberos.induct, analz_mono_contra) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1508 | apply (frule_tac [5] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1509 | apply (frule_tac [7] Says_ticket_in_parts_spies) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1510 | apply (simp_all (no_asm_simp)) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1511 | apply clarify | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1512 | txt{*K5*}
 | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1513 | apply auto | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1514 | apply (simp add: takeWhile_tail) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1515 | txt{*Level 15: case study necessary because the assumption doesn't state
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1516 | the form of ServTicket. The guarantee becomes stronger.*} | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1517 | apply (blast dest: Says_imp_spies [THEN analz.Inj, THEN analz_Decrypt'] | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1518 | K3_imp_K2 K4_trustworthy' | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 1519 | parts_spies_takeWhile_mono [THEN subsetD] | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1520 | parts_spies_evs_revD2 [THEN subsetD] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1521 | intro: Says_Auth) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1522 | apply (simp add: takeWhile_tail) | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1523 | done | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1524 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1525 | lemma A_Knows_A_Knows_ServKey: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1526 |      "[| Says A B {|ServTicket, Crypt ServKey {|Agent A, Number Ta|}|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1527 | \<in> set evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1528 |          Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1529 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1530 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1531 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1532 | ~ ExpirAuth Tk evs; ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1533 | B \<noteq> Tgs; A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1534 |    ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
 | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1535 | by (blast dest!: Confidentiality_Serv_A A_Knows_A_Knows_ServKey_lemma) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1536 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1537 | lemma B_Knows_A_Knows_ServKey: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1538 |      "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1539 |          Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1540 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1541 |          Crypt AuthKey {|Key ServKey, Agent B, Number Tt, ServTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1542 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1543 |          Crypt (shrK A) {|Key AuthKey, Agent Tgs, Number Tk, AuthTicket|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1544 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1545 | ~ ExpirServ Tt evs; ~ ExpirAuth Tk evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1546 | B \<noteq> Tgs; A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1547 |    ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
 | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1548 | by (blast dest: A_Authenticity Confidentiality_B A_Knows_A_Knows_ServKey_lemma) | 
| 14182 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1549 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1550 | |
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1551 | lemma B_Knows_A_Knows_ServKey_refined: | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1552 |      "[| Crypt ServKey {|Agent A, Number Ta|} \<in> parts (spies evs);
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1553 |          Crypt (shrK B) {|Agent A, Agent B, Key ServKey, Number Tt|}
 | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1554 | \<in> parts (spies evs); | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1555 | ~ ExpirServ Tt evs; | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1556 | B \<noteq> Tgs; A \<notin> bad; B \<notin> bad; evs \<in> kerberos |] | 
| 
5f49f00fe084
conversion of HOL/Auth/KerberosIV to new-style theory
 paulson parents: 
13507diff
changeset | 1557 |    ==> A Issues B with (Crypt ServKey {|Agent A, Number Ta|}) on evs"
 | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14182diff
changeset | 1558 | by (blast dest: A_Authenticity_refined Confidentiality_B_refined A_Knows_A_Knows_ServKey_lemma) | 
| 6452 | 1559 | |
| 1560 | end |