| 13508 |      1 | (******************************************************************************
 | 
|  |      2 | from G. Karjoth, N. Asokan and C. Gulcu
 | 
|  |      3 | "Protecting the computation results of free-roaming agents"
 | 
|  |      4 | Mobiles Agents 1998, LNCS 1477
 | 
|  |      5 | 
 | 
|  |      6 | date: february 2002
 | 
|  |      7 | author: Frederic Blanqui
 | 
|  |      8 | email: blanqui@lri.fr
 | 
|  |      9 | webpage: http://www.lri.fr/~blanqui/
 | 
|  |     10 | 
 | 
|  |     11 | University of Cambridge, Computer Laboratory
 | 
|  |     12 | William Gates Building, JJ Thomson Avenue
 | 
|  |     13 | Cambridge CB3 0FD, United Kingdom
 | 
|  |     14 | ******************************************************************************)
 | 
|  |     15 | 
 | 
|  |     16 | header{*Protocol P1*}
 | 
|  |     17 | 
 | 
| 16417 |     18 | theory P1 imports Guard_Public List_Msg begin
 | 
| 13508 |     19 | 
 | 
|  |     20 | subsection{*Protocol Definition*}
 | 
|  |     21 | 
 | 
|  |     22 | (******************************************************************************
 | 
|  |     23 | 
 | 
|  |     24 | the contents of the messages are not completely specified in the paper
 | 
|  |     25 | we assume that the user sends his request and his itinerary in the clear
 | 
|  |     26 | 
 | 
|  |     27 | we will adopt the following format for messages: {|A,r,I,L|}
 | 
|  |     28 | A: originator (agent)
 | 
|  |     29 | r: request (number)
 | 
|  |     30 | I: next shops (agent list)
 | 
|  |     31 | L: collected offers (offer list)
 | 
|  |     32 | 
 | 
|  |     33 | in the paper, the authors use nonces r_i to add redundancy in the offer
 | 
|  |     34 | in order to make it safer against dictionary attacks
 | 
|  |     35 | it is not necessary in our modelization since crypto is assumed to be strong
 | 
|  |     36 | (Crypt in injective)
 | 
|  |     37 | ******************************************************************************)
 | 
|  |     38 | 
 | 
|  |     39 | subsubsection{*offer chaining:
 | 
|  |     40 | B chains his offer for A with the head offer of L for sending it to C*}
 | 
|  |     41 | 
 | 
|  |     42 | constdefs chain :: "agent => nat => agent => msg => agent => msg"
 | 
|  |     43 | "chain B ofr A L C ==
 | 
|  |     44 | let m1= Crypt (pubK A) (Nonce ofr) in
 | 
|  |     45 | let m2= Hash {|head L, Agent C|} in
 | 
|  |     46 | sign B {|m1,m2|}"
 | 
|  |     47 | 
 | 
|  |     48 | declare Let_def [simp]
 | 
|  |     49 | 
 | 
|  |     50 | lemma chain_inj [iff]: "(chain B ofr A L C = chain B' ofr' A' L' C')
 | 
|  |     51 | = (B=B' & ofr=ofr' & A=A' & head L = head L' & C=C')"
 | 
|  |     52 | by (auto simp: chain_def Let_def)
 | 
|  |     53 | 
 | 
|  |     54 | lemma Nonce_in_chain [iff]: "Nonce ofr:parts {chain B ofr A L C}"
 | 
|  |     55 | by (auto simp: chain_def sign_def)
 | 
|  |     56 | 
 | 
|  |     57 | subsubsection{*agent whose key is used to sign an offer*}
 | 
|  |     58 | 
 | 
|  |     59 | consts shop :: "msg => msg"
 | 
|  |     60 | 
 | 
|  |     61 | recdef shop "measure size"
 | 
|  |     62 | "shop {|B,X,Crypt K H|} = Agent (agt K)"
 | 
|  |     63 | 
 | 
|  |     64 | lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B"
 | 
|  |     65 | by (simp add: chain_def sign_def)
 | 
|  |     66 | 
 | 
|  |     67 | subsubsection{*nonce used in an offer*}
 | 
|  |     68 | 
 | 
|  |     69 | consts nonce :: "msg => msg"
 | 
|  |     70 | 
 | 
|  |     71 | recdef nonce "measure size"
 | 
|  |     72 | "nonce {|B,{|Crypt K ofr,m2|},CryptH|} = ofr"
 | 
|  |     73 | 
 | 
|  |     74 | lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr"
 | 
|  |     75 | by (simp add: chain_def sign_def)
 | 
|  |     76 | 
 | 
|  |     77 | subsubsection{*next shop*}
 | 
|  |     78 | 
 | 
|  |     79 | consts next_shop :: "msg => agent"
 | 
|  |     80 | 
 | 
|  |     81 | recdef next_shop "measure size"
 | 
|  |     82 | "next_shop {|B,{|m1,Hash{|headL,Agent C|}|},CryptH|} = C"
 | 
|  |     83 | 
 | 
|  |     84 | lemma next_shop_chain [iff]: "next_shop (chain B ofr A L C) = C"
 | 
|  |     85 | by (simp add: chain_def sign_def)
 | 
|  |     86 | 
 | 
|  |     87 | subsubsection{*anchor of the offer list*}
 | 
|  |     88 | 
 | 
|  |     89 | constdefs anchor :: "agent => nat => agent => msg"
 | 
|  |     90 | "anchor A n B == chain A n A (cons nil nil) B"
 | 
|  |     91 | 
 | 
|  |     92 | lemma anchor_inj [iff]: "(anchor A n B = anchor A' n' B')
 | 
|  |     93 | = (A=A' & n=n' & B=B')"
 | 
|  |     94 | by (auto simp: anchor_def)
 | 
|  |     95 | 
 | 
|  |     96 | lemma Nonce_in_anchor [iff]: "Nonce n:parts {anchor A n B}"
 | 
|  |     97 | by (auto simp: anchor_def)
 | 
|  |     98 | 
 | 
|  |     99 | lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A"
 | 
|  |    100 | by (simp add: anchor_def)
 | 
|  |    101 | 
 | 
|  |    102 | lemma nonce_anchor [simp]: "nonce (anchor A n B) = Nonce n"
 | 
|  |    103 | by (simp add: anchor_def)
 | 
|  |    104 | 
 | 
|  |    105 | lemma next_shop_anchor [iff]: "next_shop (anchor A n B) = B"
 | 
|  |    106 | by (simp add: anchor_def)
 | 
|  |    107 | 
 | 
|  |    108 | subsubsection{*request event*}
 | 
|  |    109 | 
 | 
|  |    110 | constdefs reqm :: "agent => nat => nat => msg => agent => msg"
 | 
|  |    111 | "reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I),
 | 
|  |    112 | cons (anchor A n B) nil|}"
 | 
|  |    113 | 
 | 
|  |    114 | lemma reqm_inj [iff]: "(reqm A r n I B = reqm A' r' n' I' B')
 | 
|  |    115 | = (A=A' & r=r' & n=n' & I=I' & B=B')"
 | 
|  |    116 | by (auto simp: reqm_def)
 | 
|  |    117 | 
 | 
|  |    118 | lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}"
 | 
|  |    119 | by (auto simp: reqm_def)
 | 
|  |    120 | 
 | 
|  |    121 | constdefs req :: "agent => nat => nat => msg => agent => event"
 | 
|  |    122 | "req A r n I B == Says A B (reqm A r n I B)"
 | 
|  |    123 | 
 | 
|  |    124 | lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B')
 | 
|  |    125 | = (A=A' & r=r' & n=n' & I=I' & B=B')"
 | 
|  |    126 | by (auto simp: req_def)
 | 
|  |    127 | 
 | 
|  |    128 | subsubsection{*propose event*}
 | 
|  |    129 | 
 | 
|  |    130 | constdefs prom :: "agent => nat => agent => nat => msg => msg =>
 | 
|  |    131 | msg => agent => msg"
 | 
|  |    132 | "prom B ofr A r I L J C == {|Agent A, Number r,
 | 
|  |    133 | app (J, del (Agent B, I)), cons (chain B ofr A L C) L|}"
 | 
|  |    134 | 
 | 
|  |    135 | lemma prom_inj [dest]: "prom B ofr A r I L J C
 | 
|  |    136 | = prom B' ofr' A' r' I' L' J' C'
 | 
|  |    137 | ==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 | 
|  |    138 | by (auto simp: prom_def)
 | 
|  |    139 | 
 | 
|  |    140 | lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}"
 | 
|  |    141 | by (auto simp: prom_def)
 | 
|  |    142 | 
 | 
|  |    143 | constdefs pro :: "agent => nat => agent => nat => msg => msg =>
 | 
|  |    144 | msg => agent => event"
 | 
|  |    145 | "pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)"
 | 
|  |    146 | 
 | 
|  |    147 | lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C'
 | 
|  |    148 | ==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
 | 
|  |    149 | by (auto simp: pro_def dest: prom_inj)
 | 
|  |    150 | 
 | 
|  |    151 | subsubsection{*protocol*}
 | 
|  |    152 | 
 | 
| 23746 |    153 | inductive_set p1 :: "event list set"
 | 
|  |    154 | where
 | 
| 13508 |    155 | 
 | 
| 23746 |    156 |   Nil: "[]:p1"
 | 
| 13508 |    157 | 
 | 
| 23746 |    158 | | Fake: "[| evsf:p1; X:synth (analz (spies evsf)) |] ==> Says Spy B X # evsf : p1"
 | 
| 13508 |    159 | 
 | 
| 23746 |    160 | | Request: "[| evsr:p1; Nonce n ~:used evsr; I:agl |] ==> req A r n I B # evsr : p1"
 | 
| 13508 |    161 | 
 | 
| 23746 |    162 | | Propose: "[| evsp:p1; Says A' B {|Agent A,Number r,I,cons M L|}:set evsp;
 | 
|  |    163 |   I:agl; J:agl; isin (Agent C, app (J, del (Agent B, I)));
 | 
|  |    164 |   Nonce ofr ~:used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp : p1"
 | 
| 13508 |    165 | 
 | 
|  |    166 | subsubsection{*Composition of Traces*}
 | 
|  |    167 | 
 | 
|  |    168 | lemma "evs':p1 ==> 
 | 
|  |    169 |        evs:p1 & (ALL n. Nonce n:used evs' --> Nonce n ~:used evs) --> 
 | 
|  |    170 |        evs'@evs : p1"
 | 
|  |    171 | apply (erule p1.induct, safe) 
 | 
|  |    172 | apply (simp_all add: used_ConsI) 
 | 
|  |    173 | apply (erule p1.Fake, erule synth_sub, rule analz_mono, rule knows_sub_app)
 | 
|  |    174 | apply (erule p1.Request, safe, simp_all add: req_def, force) 
 | 
|  |    175 | apply (erule_tac A'=A' in p1.Propose, simp_all) 
 | 
|  |    176 | apply (drule_tac x=ofr in spec, simp add: pro_def, blast) 
 | 
|  |    177 | apply (erule_tac A'=A' in p1.Propose, auto simp: pro_def)
 | 
|  |    178 | done
 | 
|  |    179 | 
 | 
|  |    180 | subsubsection{*Valid Offer Lists*}
 | 
|  |    181 | 
 | 
| 23746 |    182 | inductive_set
 | 
|  |    183 |   valid :: "agent => nat => agent => msg set"
 | 
|  |    184 |   for A :: agent and n :: nat and B :: agent
 | 
|  |    185 | where
 | 
|  |    186 |   Request [intro]: "cons (anchor A n B) nil:valid A n B"
 | 
| 13508 |    187 | 
 | 
| 23746 |    188 | | Propose [intro]: "L:valid A n B
 | 
| 13508 |    189 | ==> cons (chain (next_shop (head L)) ofr A L C) L:valid A n B"
 | 
|  |    190 | 
 | 
|  |    191 | subsubsection{*basic properties of valid*}
 | 
|  |    192 | 
 | 
|  |    193 | lemma valid_not_empty: "L:valid A n B ==> EX M L'. L = cons M L'"
 | 
|  |    194 | by (erule valid.cases, auto)
 | 
|  |    195 | 
 | 
|  |    196 | lemma valid_pos_len: "L:valid A n B ==> 0 < len L"
 | 
|  |    197 | by (erule valid.induct, auto)
 | 
|  |    198 | 
 | 
|  |    199 | subsubsection{*offers of an offer list*}
 | 
|  |    200 | 
 | 
|  |    201 | constdefs offer_nonces :: "msg => msg set"
 | 
|  |    202 | "offer_nonces L == {X. X:parts {L} & (EX n. X = Nonce n)}"
 | 
|  |    203 | 
 | 
|  |    204 | subsubsection{*the originator can get the offers*}
 | 
|  |    205 | 
 | 
|  |    206 | lemma "L:valid A n B ==> offer_nonces L <= analz (insert L (initState A))"
 | 
|  |    207 | by (erule valid.induct, auto simp: anchor_def chain_def sign_def
 | 
|  |    208 | offer_nonces_def initState.simps)
 | 
|  |    209 | 
 | 
|  |    210 | subsubsection{*list of offers*}
 | 
|  |    211 | 
 | 
|  |    212 | consts offers :: "msg => msg"
 | 
|  |    213 | 
 | 
|  |    214 | recdef offers "measure size"
 | 
|  |    215 | "offers (cons M L) = cons {|shop M, nonce M|} (offers L)"
 | 
|  |    216 | "offers other = nil"
 | 
|  |    217 | 
 | 
|  |    218 | subsubsection{*list of agents whose keys are used to sign a list of offers*}
 | 
|  |    219 | 
 | 
|  |    220 | consts shops :: "msg => msg"
 | 
|  |    221 | 
 | 
|  |    222 | recdef shops "measure size"
 | 
|  |    223 | "shops (cons M L) = cons (shop M) (shops L)"
 | 
|  |    224 | "shops other = other"
 | 
|  |    225 | 
 | 
|  |    226 | lemma shops_in_agl: "L:valid A n B ==> shops L:agl"
 | 
|  |    227 | by (erule valid.induct, auto simp: anchor_def chain_def sign_def)
 | 
|  |    228 | 
 | 
|  |    229 | subsubsection{*builds a trace from an itinerary*}
 | 
|  |    230 | 
 | 
|  |    231 | consts offer_list :: "agent * nat * agent * msg * nat => msg"
 | 
|  |    232 | 
 | 
|  |    233 | recdef offer_list "measure (%(A,n,B,I,ofr). size I)"
 | 
|  |    234 | "offer_list (A,n,B,nil,ofr) = cons (anchor A n B) nil"
 | 
|  |    235 | "offer_list (A,n,B,cons (Agent C) I,ofr) = (
 | 
|  |    236 | let L = offer_list (A,n,B,I,Suc ofr) in
 | 
|  |    237 | cons (chain (next_shop (head L)) ofr A L C) L)"
 | 
|  |    238 | 
 | 
|  |    239 | lemma "I:agl ==> ALL ofr. offer_list (A,n,B,I,ofr):valid A n B"
 | 
|  |    240 | by (erule agl.induct, auto)
 | 
|  |    241 | 
 | 
|  |    242 | consts trace :: "agent * nat * agent * nat * msg * msg * msg
 | 
|  |    243 | => event list"
 | 
|  |    244 | 
 | 
|  |    245 | recdef trace "measure (%(B,ofr,A,r,I,L,K). size K)"
 | 
|  |    246 | "trace (B,ofr,A,r,I,L,nil) = []"
 | 
|  |    247 | "trace (B,ofr,A,r,I,L,cons (Agent D) K) = (
 | 
|  |    248 | let C = (if K=nil then B else agt_nb (head K)) in
 | 
|  |    249 | let I' = (if K=nil then cons (Agent A) (cons (Agent B) I)
 | 
|  |    250 |           else cons (Agent A) (app (I, cons (head K) nil))) in
 | 
|  |    251 | let I'' = app (I, cons (head K) nil) in
 | 
|  |    252 | pro C (Suc ofr) A r I' L nil D
 | 
|  |    253 | # trace (B,Suc ofr,A,r,I'',tail L,K))"
 | 
|  |    254 | 
 | 
|  |    255 | constdefs trace' :: "agent => nat => nat => msg => agent => nat => event list"
 | 
|  |    256 | "trace' A r n I B ofr == (
 | 
|  |    257 | let AI = cons (Agent A) I in
 | 
|  |    258 | let L = offer_list (A,n,B,AI,ofr) in
 | 
|  |    259 | trace (B,ofr,A,r,nil,L,AI))"
 | 
|  |    260 | 
 | 
|  |    261 | declare trace'_def [simp]
 | 
|  |    262 | 
 | 
|  |    263 | subsubsection{*there is a trace in which the originator receives a valid answer*}
 | 
|  |    264 | 
 | 
|  |    265 | lemma p1_not_empty: "evs:p1 ==> req A r n I B:set evs -->
 | 
|  |    266 | (EX evs'. evs'@evs:p1 & pro B' ofr A r I' L J A:set evs' & L:valid A n B)"
 | 
|  |    267 | oops
 | 
|  |    268 | 
 | 
|  |    269 | 
 | 
|  |    270 | subsection{*properties of protocol P1*}
 | 
|  |    271 | 
 | 
|  |    272 | text{*publicly verifiable forward integrity:
 | 
|  |    273 | anyone can verify the validity of an offer list*}
 | 
|  |    274 | 
 | 
|  |    275 | subsubsection{*strong forward integrity:
 | 
|  |    276 | except the last one, no offer can be modified*}
 | 
|  |    277 | 
 | 
|  |    278 | lemma strong_forward_integrity: "ALL L. Suc i < len L
 | 
|  |    279 | --> L:valid A n B & repl (L,Suc i,M):valid A n B --> M = ith (L,Suc i)"
 | 
|  |    280 | apply (induct i)
 | 
|  |    281 | (* i = 0 *)
 | 
|  |    282 | apply clarify
 | 
|  |    283 | apply (frule len_not_empty, clarsimp)
 | 
|  |    284 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    285 | apply (ind_cases "{|x,xa,l'a|}:valid A n B" for x xa l'a)
 | 
|  |    286 | apply (ind_cases "{|x,M,l'a|}:valid A n B" for x l'a)
 | 
| 13508 |    287 | apply (simp add: chain_def)
 | 
|  |    288 | (* i > 0 *)
 | 
|  |    289 | apply clarify
 | 
|  |    290 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    291 | apply (ind_cases "{|x,repl(l',Suc na,M)|}:valid A n B" for x l' na)
 | 
| 13508 |    292 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    293 | apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 13508 |    294 | by (drule_tac x=l' in spec, simp, blast)
 | 
|  |    295 | 
 | 
|  |    296 | subsubsection{*insertion resilience:
 | 
|  |    297 | except at the beginning, no offer can be inserted*}
 | 
|  |    298 | 
 | 
|  |    299 | lemma chain_isnt_head [simp]: "L:valid A n B ==>
 | 
|  |    300 | head L ~= chain (next_shop (head L)) ofr A L C"
 | 
|  |    301 | by (erule valid.induct, auto simp: chain_def sign_def anchor_def)
 | 
|  |    302 | 
 | 
|  |    303 | lemma insertion_resilience: "ALL L. L:valid A n B --> Suc i < len L
 | 
|  |    304 | --> ins (L,Suc i,M) ~:valid A n B"
 | 
|  |    305 | apply (induct i)
 | 
|  |    306 | (* i = 0 *)
 | 
|  |    307 | apply clarify
 | 
|  |    308 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    309 | apply (ind_cases "{|x,l'|}:valid A n B" for x l', simp)
 | 
|  |    310 | apply (ind_cases "{|x,M,l'|}:valid A n B" for x l', clarsimp)
 | 
|  |    311 | apply (ind_cases "{|head l',l'|}:valid A n B" for l', simp, simp)
 | 
| 13508 |    312 | (* i > 0 *)
 | 
|  |    313 | apply clarify
 | 
|  |    314 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    315 | apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 13508 |    316 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    317 | apply (ind_cases "{|x,ins(l',Suc na,M)|}:valid A n B" for x l' na)
 | 
| 13508 |    318 | apply (frule len_not_empty, clarsimp)
 | 
|  |    319 | by (drule_tac x=l' in spec, clarsimp)
 | 
|  |    320 | 
 | 
|  |    321 | subsubsection{*truncation resilience:
 | 
|  |    322 | only shop i can truncate at offer i*}
 | 
|  |    323 | 
 | 
|  |    324 | lemma truncation_resilience: "ALL L. L:valid A n B --> Suc i < len L
 | 
|  |    325 | --> cons M (trunc (L,Suc i)):valid A n B --> shop M = shop (ith (L,i))"
 | 
|  |    326 | apply (induct i)
 | 
|  |    327 | (* i = 0 *)
 | 
|  |    328 | apply clarify
 | 
|  |    329 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    330 | apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 13508 |    331 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    332 | apply (ind_cases "{|M,l'|}:valid A n B" for l')
 | 
| 13508 |    333 | apply (frule len_not_empty, clarsimp, simp)
 | 
|  |    334 | (* i > 0 *)
 | 
|  |    335 | apply clarify
 | 
|  |    336 | apply (frule len_not_empty, clarsimp)
 | 
| 23746 |    337 | apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 13508 |    338 | apply (frule len_not_empty, clarsimp)
 | 
|  |    339 | by (drule_tac x=l' in spec, clarsimp)
 | 
|  |    340 | 
 | 
|  |    341 | subsubsection{*declarations for tactics*}
 | 
|  |    342 | 
 | 
|  |    343 | declare knows_Spy_partsEs [elim]
 | 
|  |    344 | declare Fake_parts_insert [THEN subsetD, dest]
 | 
|  |    345 | declare initState.simps [simp del]
 | 
|  |    346 | 
 | 
|  |    347 | subsubsection{*get components of a message*}
 | 
|  |    348 | 
 | 
|  |    349 | lemma get_ML [dest]: "Says A' B {|A,r,I,M,L|}:set evs ==>
 | 
|  |    350 | M:parts (spies evs) & L:parts (spies evs)"
 | 
|  |    351 | by blast
 | 
|  |    352 | 
 | 
|  |    353 | subsubsection{*general properties of p1*}
 | 
|  |    354 | 
 | 
|  |    355 | lemma reqm_neq_prom [iff]:
 | 
|  |    356 | "reqm A r n I B ~= prom B' ofr A' r' I' (cons M L) J C"
 | 
|  |    357 | by (auto simp: reqm_def prom_def)
 | 
|  |    358 | 
 | 
|  |    359 | lemma prom_neq_reqm [iff]:
 | 
|  |    360 | "prom B' ofr A' r' I' (cons M L) J C ~= reqm A r n I B"
 | 
|  |    361 | by (auto simp: reqm_def prom_def)
 | 
|  |    362 | 
 | 
|  |    363 | lemma req_neq_pro [iff]: "req A r n I B ~= pro B' ofr A' r' I' (cons M L) J C"
 | 
|  |    364 | by (auto simp: req_def pro_def)
 | 
|  |    365 | 
 | 
|  |    366 | lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C ~= req A r n I B"
 | 
|  |    367 | by (auto simp: req_def pro_def)
 | 
|  |    368 | 
 | 
|  |    369 | lemma p1_has_no_Gets: "evs:p1 ==> ALL A X. Gets A X ~:set evs"
 | 
|  |    370 | by (erule p1.induct, auto simp: req_def pro_def)
 | 
|  |    371 | 
 | 
|  |    372 | lemma p1_is_Gets_correct [iff]: "Gets_correct p1"
 | 
|  |    373 | by (auto simp: Gets_correct_def dest: p1_has_no_Gets)
 | 
|  |    374 | 
 | 
|  |    375 | lemma p1_is_one_step [iff]: "one_step p1"
 | 
| 23746 |    376 | by (unfold one_step_def, clarify, ind_cases "ev#evs:p1" for ev evs, auto)
 | 
| 13508 |    377 | 
 | 
|  |    378 | lemma p1_has_only_Says' [rule_format]: "evs:p1 ==>
 | 
|  |    379 | ev:set evs --> (EX A B X. ev=Says A B X)"
 | 
|  |    380 | by (erule p1.induct, auto simp: req_def pro_def)
 | 
|  |    381 | 
 | 
|  |    382 | lemma p1_has_only_Says [iff]: "has_only_Says p1"
 | 
|  |    383 | by (auto simp: has_only_Says_def dest: p1_has_only_Says')
 | 
|  |    384 | 
 | 
|  |    385 | lemma p1_is_regular [iff]: "regular p1"
 | 
|  |    386 | apply (simp only: regular_def, clarify)
 | 
|  |    387 | apply (erule_tac p1.induct)
 | 
|  |    388 | apply (simp_all add: initState.simps knows.simps pro_def prom_def
 | 
|  |    389 |                      req_def reqm_def anchor_def chain_def sign_def)
 | 
|  |    390 | by (auto dest: no_Key_in_agl no_Key_in_appdel parts_trans)
 | 
|  |    391 | 
 | 
|  |    392 | subsubsection{*private keys are safe*}
 | 
|  |    393 | 
 | 
|  |    394 | lemma priK_parts_Friend_imp_bad [rule_format,dest]:
 | 
|  |    395 |      "[| evs:p1; Friend B ~= A |]
 | 
|  |    396 |       ==> (Key (priK A):parts (knows (Friend B) evs)) --> (A:bad)"
 | 
|  |    397 | apply (erule p1.induct)
 | 
|  |    398 | apply (simp_all add: initState.simps knows.simps pro_def prom_def
 | 
| 17778 |    399 |                 req_def reqm_def anchor_def chain_def sign_def)
 | 
| 13508 |    400 | apply (blast dest: no_Key_in_agl)
 | 
|  |    401 | apply (auto del: parts_invKey disjE  dest: parts_trans
 | 
|  |    402 |             simp add: no_Key_in_appdel)
 | 
|  |    403 | done
 | 
|  |    404 | 
 | 
|  |    405 | lemma priK_analz_Friend_imp_bad [rule_format,dest]:
 | 
|  |    406 |      "[| evs:p1; Friend B ~= A |]
 | 
|  |    407 | ==> (Key (priK A):analz (knows (Friend B) evs)) --> (A:bad)"
 | 
|  |    408 | by auto
 | 
|  |    409 | 
 | 
|  |    410 | lemma priK_notin_knows_max_Friend: "[| evs:p1; A ~:bad; A ~= Friend C |]
 | 
|  |    411 | ==> Key (priK A) ~:analz (knows_max (Friend C) evs)"
 | 
|  |    412 | apply (rule not_parts_not_analz, simp add: knows_max_def, safe)
 | 
|  |    413 | apply (drule_tac H="spies' evs" in parts_sub)
 | 
|  |    414 | apply (rule_tac p=p1 in knows_max'_sub_spies', simp+)
 | 
|  |    415 | apply (drule_tac H="spies evs" in parts_sub)
 | 
|  |    416 | by (auto dest: knows'_sub_knows [THEN subsetD] priK_notin_initState_Friend)
 | 
|  |    417 | 
 | 
|  |    418 | subsubsection{*general guardedness properties*}
 | 
|  |    419 | 
 | 
|  |    420 | lemma agl_guard [intro]: "I:agl ==> I:guard n Ks"
 | 
|  |    421 | by (erule agl.induct, auto)
 | 
|  |    422 | 
 | 
|  |    423 | lemma Says_to_knows_max'_guard: "[| Says A' C {|A'',r,I,L|}:set evs;
 | 
|  |    424 | Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks"
 | 
|  |    425 | by (auto dest: Says_to_knows_max')
 | 
|  |    426 | 
 | 
|  |    427 | lemma Says_from_knows_max'_guard: "[| Says C A' {|A'',r,I,L|}:set evs;
 | 
|  |    428 | Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks"
 | 
|  |    429 | by (auto dest: Says_from_knows_max')
 | 
|  |    430 | 
 | 
|  |    431 | lemma Says_Nonce_not_used_guard: "[| Says A' B {|A'',r,I,L|}:set evs;
 | 
|  |    432 | Nonce n ~:used evs |] ==> L:guard n Ks"
 | 
|  |    433 | by (drule not_used_not_parts, auto)
 | 
|  |    434 | 
 | 
|  |    435 | subsubsection{*guardedness of messages*}
 | 
|  |    436 | 
 | 
|  |    437 | lemma chain_guard [iff]: "chain B ofr A L C:guard n {priK A}"
 | 
|  |    438 | by (case_tac "ofr=n", auto simp: chain_def sign_def)
 | 
|  |    439 | 
 | 
|  |    440 | lemma chain_guard_Nonce_neq [intro]: "n ~= ofr
 | 
|  |    441 | ==> chain B ofr A' L C:guard n {priK A}"
 | 
|  |    442 | by (auto simp: chain_def sign_def)
 | 
|  |    443 | 
 | 
|  |    444 | lemma anchor_guard [iff]: "anchor A n' B:guard n {priK A}"
 | 
|  |    445 | by (case_tac "n'=n", auto simp: anchor_def)
 | 
|  |    446 | 
 | 
|  |    447 | lemma anchor_guard_Nonce_neq [intro]: "n ~= n'
 | 
|  |    448 | ==> anchor A' n' B:guard n {priK A}"
 | 
|  |    449 | by (auto simp: anchor_def)
 | 
|  |    450 | 
 | 
|  |    451 | lemma reqm_guard [intro]: "I:agl ==> reqm A r n' I B:guard n {priK A}"
 | 
|  |    452 | by (case_tac "n'=n", auto simp: reqm_def)
 | 
|  |    453 | 
 | 
|  |    454 | lemma reqm_guard_Nonce_neq [intro]: "[| n ~= n'; I:agl |]
 | 
|  |    455 | ==> reqm A' r n' I B:guard n {priK A}"
 | 
|  |    456 | by (auto simp: reqm_def)
 | 
|  |    457 | 
 | 
|  |    458 | lemma prom_guard [intro]: "[| I:agl; J:agl; L:guard n {priK A} |]
 | 
|  |    459 | ==> prom B ofr A r I L J C:guard n {priK A}"
 | 
|  |    460 | by (auto simp: prom_def)
 | 
|  |    461 | 
 | 
|  |    462 | lemma prom_guard_Nonce_neq [intro]: "[| n ~= ofr; I:agl; J:agl;
 | 
|  |    463 | L:guard n {priK A} |] ==> prom B ofr A' r I L J C:guard n {priK A}"
 | 
|  |    464 | by (auto simp: prom_def)
 | 
|  |    465 | 
 | 
|  |    466 | subsubsection{*Nonce uniqueness*}
 | 
|  |    467 | 
 | 
|  |    468 | lemma uniq_Nonce_in_chain [dest]: "Nonce k:parts {chain B ofr A L C} ==> k=ofr"
 | 
|  |    469 | by (auto simp: chain_def sign_def)
 | 
|  |    470 | 
 | 
|  |    471 | lemma uniq_Nonce_in_anchor [dest]: "Nonce k:parts {anchor A n B} ==> k=n"
 | 
|  |    472 | by (auto simp: anchor_def chain_def sign_def)
 | 
|  |    473 | 
 | 
|  |    474 | lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k:parts {reqm A r n I B};
 | 
|  |    475 | I:agl |] ==> k=n"
 | 
|  |    476 | by (auto simp: reqm_def dest: no_Nonce_in_agl)
 | 
|  |    477 | 
 | 
|  |    478 | lemma uniq_Nonce_in_prom [dest]: "[| Nonce k:parts {prom B ofr A r I L J C};
 | 
|  |    479 | I:agl; J:agl; Nonce k ~:parts {L} |] ==> k=ofr"
 | 
|  |    480 | by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel)
 | 
|  |    481 | 
 | 
|  |    482 | subsubsection{*requests are guarded*}
 | 
|  |    483 | 
 | 
|  |    484 | lemma req_imp_Guard [rule_format]: "[| evs:p1; A ~:bad |] ==>
 | 
|  |    485 | req A r n I B:set evs --> Guard n {priK A} (spies evs)"
 | 
|  |    486 | apply (erule p1.induct, simp)
 | 
|  |    487 | apply (simp add: req_def knows.simps, safe)
 | 
|  |    488 | apply (erule in_synth_Guard, erule Guard_analz, simp)
 | 
|  |    489 | by (auto simp: req_def pro_def dest: Says_imp_knows_Spy)
 | 
|  |    490 | 
 | 
|  |    491 | lemma req_imp_Guard_Friend: "[| evs:p1; A ~:bad; req A r n I B:set evs |]
 | 
|  |    492 | ==> Guard n {priK A} (knows_max (Friend C) evs)"
 | 
|  |    493 | apply (rule Guard_knows_max')
 | 
|  |    494 | apply (rule_tac H="spies evs" in Guard_mono)
 | 
|  |    495 | apply (rule req_imp_Guard, simp+)
 | 
|  |    496 | apply (rule_tac B="spies' evs" in subset_trans)
 | 
|  |    497 | apply (rule_tac p=p1 in knows_max'_sub_spies', simp+)
 | 
|  |    498 | by (rule knows'_sub_knows)
 | 
|  |    499 | 
 | 
|  |    500 | subsubsection{*propositions are guarded*}
 | 
|  |    501 | 
 | 
|  |    502 | lemma pro_imp_Guard [rule_format]: "[| evs:p1; B ~:bad; A ~:bad |] ==>
 | 
|  |    503 | pro B ofr A r I (cons M L) J C:set evs --> Guard ofr {priK A} (spies evs)"
 | 
|  |    504 | apply (erule p1.induct) (* +3 subgoals *)
 | 
|  |    505 | (* Nil *)
 | 
|  |    506 | apply simp
 | 
|  |    507 | (* Fake *)
 | 
|  |    508 | apply (simp add: pro_def, safe) (* +4 subgoals *)
 | 
|  |    509 | (* 1 *)
 | 
|  |    510 | apply (erule in_synth_Guard, drule Guard_analz, simp, simp)
 | 
|  |    511 | (* 2 *)
 | 
|  |    512 | apply simp
 | 
|  |    513 | (* 3 *)
 | 
|  |    514 | apply (simp, simp add: req_def pro_def, blast)
 | 
|  |    515 | (* 4 *)
 | 
|  |    516 | apply (simp add: pro_def)
 | 
|  |    517 | apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard)
 | 
|  |    518 | (* 5 *)
 | 
|  |    519 | apply simp
 | 
|  |    520 | apply safe (* +1 subgoal *)
 | 
|  |    521 | apply (simp add: pro_def)
 | 
|  |    522 | apply (blast dest: prom_inj Says_Nonce_not_used_guard)
 | 
|  |    523 | (* 6 *)
 | 
|  |    524 | apply (simp add: pro_def)
 | 
|  |    525 | apply (blast dest: Says_imp_knows_Spy)
 | 
|  |    526 | (* Request *)
 | 
|  |    527 | apply (simp add: pro_def)
 | 
|  |    528 | apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard)
 | 
|  |    529 | (* Propose *)
 | 
|  |    530 | apply simp
 | 
|  |    531 | apply safe (* +1 subgoal *)
 | 
|  |    532 | (* 1 *)
 | 
|  |    533 | apply (simp add: pro_def)
 | 
|  |    534 | apply (blast dest: prom_inj Says_Nonce_not_used_guard)
 | 
|  |    535 | (* 2 *)
 | 
|  |    536 | apply (simp add: pro_def)
 | 
|  |    537 | by (blast dest: Says_imp_knows_Spy)
 | 
|  |    538 | 
 | 
|  |    539 | lemma pro_imp_Guard_Friend: "[| evs:p1; B ~:bad; A ~:bad;
 | 
|  |    540 | pro B ofr A r I (cons M L) J C:set evs |]
 | 
|  |    541 | ==> Guard ofr {priK A} (knows_max (Friend D) evs)"
 | 
|  |    542 | apply (rule Guard_knows_max')
 | 
|  |    543 | apply (rule_tac H="spies evs" in Guard_mono)
 | 
|  |    544 | apply (rule pro_imp_Guard, simp+)
 | 
|  |    545 | apply (rule_tac B="spies' evs" in subset_trans)
 | 
|  |    546 | apply (rule_tac p=p1 in knows_max'_sub_spies', simp+)
 | 
|  |    547 | by (rule knows'_sub_knows)
 | 
|  |    548 | 
 | 
|  |    549 | subsubsection{*data confidentiality:
 | 
|  |    550 | no one other than the originator can decrypt the offers*}
 | 
|  |    551 | 
 | 
|  |    552 | lemma Nonce_req_notin_spies: "[| evs:p1; req A r n I B:set evs; A ~:bad |]
 | 
|  |    553 | ==> Nonce n ~:analz (spies evs)"
 | 
|  |    554 | by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
 | 
|  |    555 | 
 | 
|  |    556 | lemma Nonce_req_notin_knows_max_Friend: "[| evs:p1; req A r n I B:set evs;
 | 
|  |    557 | A ~:bad; A ~= Friend C |] ==> Nonce n ~:analz (knows_max (Friend C) evs)"
 | 
|  |    558 | apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+)
 | 
|  |    559 | apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+)
 | 
|  |    560 | by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def)
 | 
|  |    561 | 
 | 
|  |    562 | lemma Nonce_pro_notin_spies: "[| evs:p1; B ~:bad; A ~:bad;
 | 
|  |    563 | pro B ofr A r I (cons M L) J C:set evs |] ==> Nonce ofr ~:analz (spies evs)"
 | 
|  |    564 | by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
 | 
|  |    565 | 
 | 
|  |    566 | lemma Nonce_pro_notin_knows_max_Friend: "[| evs:p1; B ~:bad; A ~:bad;
 | 
|  |    567 | A ~= Friend D; pro B ofr A r I (cons M L) J C:set evs |]
 | 
|  |    568 | ==> Nonce ofr ~:analz (knows_max (Friend D) evs)"
 | 
|  |    569 | apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+)
 | 
|  |    570 | apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+)
 | 
|  |    571 | by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def)
 | 
|  |    572 | 
 | 
|  |    573 | subsubsection{*non repudiability:
 | 
|  |    574 | an offer signed by B has been sent by B*}
 | 
|  |    575 | 
 | 
|  |    576 | lemma Crypt_reqm: "[| Crypt (priK A) X:parts {reqm A' r n I B}; I:agl |] ==> A=A'"
 | 
|  |    577 | by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl)
 | 
|  |    578 | 
 | 
|  |    579 | lemma Crypt_prom: "[| Crypt (priK A) X:parts {prom B ofr A' r I L J C};
 | 
|  |    580 | I:agl; J:agl |] ==> A=B | Crypt (priK A) X:parts {L}"
 | 
|  |    581 | apply (simp add: prom_def anchor_def chain_def sign_def)
 | 
|  |    582 | by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel)
 | 
|  |    583 | 
 | 
|  |    584 | lemma Crypt_safeness: "[| evs:p1; A ~:bad |] ==> Crypt (priK A) X:parts (spies evs)
 | 
|  |    585 | --> (EX B Y. Says A B Y:set evs & Crypt (priK A) X:parts {Y})"
 | 
|  |    586 | apply (erule p1.induct)
 | 
|  |    587 | (* Nil *)
 | 
|  |    588 | apply simp
 | 
|  |    589 | (* Fake *)
 | 
|  |    590 | apply clarsimp
 | 
|  |    591 | apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp)
 | 
|  |    592 | apply (erule disjE)
 | 
|  |    593 | apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast)
 | 
|  |    594 | (* Request *)
 | 
|  |    595 | apply (simp add: req_def, clarify)
 | 
|  |    596 | apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp)
 | 
|  |    597 | apply (erule disjE)
 | 
|  |    598 | apply (frule Crypt_reqm, simp, clarify)
 | 
|  |    599 | apply (rule_tac x=B in exI, rule_tac x="reqm A r n I B" in exI, simp, blast)
 | 
|  |    600 | (* Propose *)
 | 
|  |    601 | apply (simp add: pro_def, clarify)
 | 
|  |    602 | apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp)
 | 
|  |    603 | apply (rotate_tac -1, erule disjE)
 | 
|  |    604 | apply (frule Crypt_prom, simp, simp)
 | 
|  |    605 | apply (rotate_tac -1, erule disjE)
 | 
|  |    606 | apply (rule_tac x=C in exI)
 | 
|  |    607 | apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI, blast)
 | 
|  |    608 | apply (subgoal_tac "cons M L:parts (spies evsp)")
 | 
|  |    609 | apply (drule_tac G="{cons M L}" and H="spies evsp" in parts_trans, blast, blast)
 | 
|  |    610 | apply (drule Says_imp_spies, rotate_tac -1, drule parts.Inj)
 | 
|  |    611 | apply (drule parts.Snd, drule parts.Snd, drule parts.Snd)
 | 
|  |    612 | by auto
 | 
|  |    613 | 
 | 
|  |    614 | lemma Crypt_Hash_imp_sign: "[| evs:p1; A ~:bad |] ==>
 | 
|  |    615 | Crypt (priK A) (Hash X):parts (spies evs)
 | 
|  |    616 | --> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})"
 | 
|  |    617 | apply (erule p1.induct)
 | 
|  |    618 | (* Nil *)
 | 
|  |    619 | apply simp
 | 
|  |    620 | (* Fake *)
 | 
|  |    621 | apply clarsimp
 | 
|  |    622 | apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD)
 | 
|  |    623 | apply simp
 | 
|  |    624 | apply (erule disjE)
 | 
|  |    625 | apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast)
 | 
|  |    626 | (* Request *)
 | 
|  |    627 | apply (simp add: req_def, clarify)
 | 
|  |    628 | apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD)
 | 
|  |    629 | apply simp
 | 
|  |    630 | apply (erule disjE)
 | 
|  |    631 | apply (frule Crypt_reqm, simp+)
 | 
|  |    632 | apply (rule_tac x=B in exI, rule_tac x="reqm Aa r n I B" in exI)
 | 
|  |    633 | apply (simp add: reqm_def sign_def anchor_def no_Crypt_in_agl)
 | 
|  |    634 | apply (simp add: chain_def sign_def, blast)
 | 
|  |    635 | (* Propose *)
 | 
|  |    636 | apply (simp add: pro_def, clarify)
 | 
|  |    637 | apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD)
 | 
|  |    638 | apply simp
 | 
|  |    639 | apply (rotate_tac -1, erule disjE)
 | 
|  |    640 | apply (simp add: prom_def sign_def no_Crypt_in_agl no_Crypt_in_appdel)
 | 
|  |    641 | apply (simp add: chain_def sign_def)
 | 
|  |    642 | apply (rotate_tac -1, erule disjE)
 | 
|  |    643 | apply (rule_tac x=C in exI)
 | 
|  |    644 | apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI)
 | 
|  |    645 | apply (simp add: prom_def chain_def sign_def)
 | 
|  |    646 | apply (erule impE) 
 | 
|  |    647 | apply (blast dest: get_ML parts_sub) 
 | 
|  |    648 | apply (blast del: MPair_parts)+
 | 
|  |    649 | done
 | 
|  |    650 | 
 | 
|  |    651 | lemma sign_safeness: "[| evs:p1; A ~:bad |] ==> sign A X:parts (spies evs)
 | 
|  |    652 | --> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})"
 | 
|  |    653 | apply (clarify, simp add: sign_def, frule parts.Snd)
 | 
|  |    654 | apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def])
 | 
|  |    655 | done
 | 
|  |    656 | 
 | 
|  |    657 | end |