| 
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: march 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 P2*}
 | 
| 
 | 
    17  | 
  | 
| 
16417
 | 
    18  | 
theory P2 imports Guard_Public List_Msg begin
  | 
| 
13508
 | 
    19  | 
  | 
| 
 | 
    20  | 
subsection{*Protocol Definition*}
 | 
| 
 | 
    21  | 
  | 
| 
 | 
    22  | 
  | 
| 
 | 
    23  | 
text{*Like P1 except the definitions of @{text chain}, @{text shop},
 | 
| 
 | 
    24  | 
  @{text next_shop} and @{text nonce}*}
 | 
| 
 | 
    25  | 
  | 
| 
 | 
    26  | 
subsubsection{*offer chaining:
 | 
| 
 | 
    27  | 
B chains his offer for A with the head offer of L for sending it to C*}
  | 
| 
 | 
    28  | 
  | 
| 
 | 
    29  | 
constdefs chain :: "agent => nat => agent => msg => agent => msg"
  | 
| 
 | 
    30  | 
"chain B ofr A L C ==
  | 
| 
 | 
    31  | 
let m1= sign B (Nonce ofr) in
  | 
| 
 | 
    32  | 
let m2= Hash {|head L, Agent C|} in
 | 
| 
 | 
    33  | 
{|Crypt (pubK A) m1, m2|}"
 | 
| 
 | 
    34  | 
  | 
| 
 | 
    35  | 
declare Let_def [simp]
  | 
| 
 | 
    36  | 
  | 
| 
 | 
    37  | 
lemma chain_inj [iff]: "(chain B ofr A L C = chain B' ofr' A' L' C')
  | 
| 
 | 
    38  | 
= (B=B' & ofr=ofr' & A=A' & head L = head L' & C=C')"
  | 
| 
 | 
    39  | 
by (auto simp: chain_def Let_def)
  | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
lemma Nonce_in_chain [iff]: "Nonce ofr:parts {chain B ofr A L C}"
 | 
| 
 | 
    42  | 
by (auto simp: chain_def sign_def)
  | 
| 
 | 
    43  | 
  | 
| 
 | 
    44  | 
subsubsection{*agent whose key is used to sign an offer*}
 | 
| 
 | 
    45  | 
  | 
| 
 | 
    46  | 
consts shop :: "msg => msg"
  | 
| 
 | 
    47  | 
  | 
| 
 | 
    48  | 
recdef shop "measure size"
  | 
| 
 | 
    49  | 
"shop {|Crypt K {|B,ofr,Crypt K' H|},m2|} = Agent (agt K')"
 | 
| 
 | 
    50  | 
  | 
| 
 | 
    51  | 
lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B"
  | 
| 
 | 
    52  | 
by (simp add: chain_def sign_def)
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
subsubsection{*nonce used in an offer*}
 | 
| 
 | 
    55  | 
  | 
| 
 | 
    56  | 
consts nonce :: "msg => msg"
  | 
| 
 | 
    57  | 
  | 
| 
 | 
    58  | 
recdef nonce "measure size"
  | 
| 
 | 
    59  | 
"nonce {|Crypt K {|B,ofr,CryptH|},m2|} = ofr"
 | 
| 
 | 
    60  | 
  | 
| 
 | 
    61  | 
lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr"
  | 
| 
 | 
    62  | 
by (simp add: chain_def sign_def)
  | 
| 
 | 
    63  | 
  | 
| 
 | 
    64  | 
subsubsection{*next shop*}
 | 
| 
 | 
    65  | 
  | 
| 
 | 
    66  | 
consts next_shop :: "msg => agent"
  | 
| 
 | 
    67  | 
  | 
| 
 | 
    68  | 
recdef next_shop "measure size"
  | 
| 
 | 
    69  | 
"next_shop {|m1,Hash {|headL,Agent C|}|} = C"
 | 
| 
 | 
    70  | 
  | 
| 
 | 
    71  | 
lemma "next_shop (chain B ofr A L C) = C"
  | 
| 
 | 
    72  | 
by (simp add: chain_def sign_def)
  | 
| 
 | 
    73  | 
  | 
| 
 | 
    74  | 
subsubsection{*anchor of the offer list*}
 | 
| 
 | 
    75  | 
  | 
| 
 | 
    76  | 
constdefs anchor :: "agent => nat => agent => msg"
  | 
| 
 | 
    77  | 
"anchor A n B == chain A n A (cons nil nil) B"
  | 
| 
 | 
    78  | 
  | 
| 
 | 
    79  | 
lemma anchor_inj [iff]:
  | 
| 
 | 
    80  | 
     "(anchor A n B = anchor A' n' B') = (A=A' & n=n' & B=B')"
  | 
| 
 | 
    81  | 
by (auto simp: anchor_def)
  | 
| 
 | 
    82  | 
  | 
| 
 | 
    83  | 
lemma Nonce_in_anchor [iff]: "Nonce n:parts {anchor A n B}"
 | 
| 
 | 
    84  | 
by (auto simp: anchor_def)
  | 
| 
 | 
    85  | 
  | 
| 
 | 
    86  | 
lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A"
  | 
| 
 | 
    87  | 
by (simp add: anchor_def)
  | 
| 
 | 
    88  | 
  | 
| 
 | 
    89  | 
subsubsection{*request event*}
 | 
| 
 | 
    90  | 
  | 
| 
 | 
    91  | 
constdefs reqm :: "agent => nat => nat => msg => agent => msg"
  | 
| 
 | 
    92  | 
"reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I),
 | 
| 
 | 
    93  | 
cons (anchor A n B) nil|}"
  | 
| 
 | 
    94  | 
  | 
| 
 | 
    95  | 
lemma reqm_inj [iff]: "(reqm A r n I B = reqm A' r' n' I' B')
  | 
| 
 | 
    96  | 
= (A=A' & r=r' & n=n' & I=I' & B=B')"
  | 
| 
 | 
    97  | 
by (auto simp: reqm_def)
  | 
| 
 | 
    98  | 
  | 
| 
 | 
    99  | 
lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}"
 | 
| 
 | 
   100  | 
by (auto simp: reqm_def)
  | 
| 
 | 
   101  | 
  | 
| 
 | 
   102  | 
constdefs req :: "agent => nat => nat => msg => agent => event"
  | 
| 
 | 
   103  | 
"req A r n I B == Says A B (reqm A r n I B)"
  | 
| 
 | 
   104  | 
  | 
| 
 | 
   105  | 
lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B')
  | 
| 
 | 
   106  | 
= (A=A' & r=r' & n=n' & I=I' & B=B')"
  | 
| 
 | 
   107  | 
by (auto simp: req_def)
  | 
| 
 | 
   108  | 
  | 
| 
 | 
   109  | 
subsubsection{*propose event*}
 | 
| 
 | 
   110  | 
  | 
| 
 | 
   111  | 
constdefs prom :: "agent => nat => agent => nat => msg => msg =>
  | 
| 
 | 
   112  | 
msg => agent => msg"
  | 
| 
 | 
   113  | 
"prom B ofr A r I L J C == {|Agent A, Number r,
 | 
| 
 | 
   114  | 
app (J, del (Agent B, I)), cons (chain B ofr A L C) L|}"
  | 
| 
 | 
   115  | 
  | 
| 
 | 
   116  | 
lemma prom_inj [dest]: "prom B ofr A r I L J C = prom B' ofr' A' r' I' L' J' C'
  | 
| 
 | 
   117  | 
==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
  | 
| 
 | 
   118  | 
by (auto simp: prom_def)
  | 
| 
 | 
   119  | 
  | 
| 
 | 
   120  | 
lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}"
 | 
| 
 | 
   121  | 
by (auto simp: prom_def)
  | 
| 
 | 
   122  | 
  | 
| 
 | 
   123  | 
constdefs pro :: "agent => nat => agent => nat => msg => msg =>
  | 
| 
 | 
   124  | 
                  msg => agent => event"
  | 
| 
 | 
   125  | 
"pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)"
  | 
| 
 | 
   126  | 
  | 
| 
 | 
   127  | 
lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C'
  | 
| 
 | 
   128  | 
==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'"
  | 
| 
 | 
   129  | 
by (auto simp: pro_def dest: prom_inj)
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
subsubsection{*protocol*}
 | 
| 
 | 
   132  | 
  | 
| 
23746
 | 
   133  | 
inductive_set p2 :: "event list set"
  | 
| 
 | 
   134  | 
where
  | 
| 
13508
 | 
   135  | 
  | 
| 
23746
 | 
   136  | 
  Nil: "[]:p2"
  | 
| 
13508
 | 
   137  | 
  | 
| 
23746
 | 
   138  | 
| Fake: "[| evsf:p2; X:synth (analz (spies evsf)) |] ==> Says Spy B X # evsf : p2"
  | 
| 
13508
 | 
   139  | 
  | 
| 
23746
 | 
   140  | 
| Request: "[| evsr:p2; Nonce n ~:used evsr; I:agl |] ==> req A r n I B # evsr : p2"
  | 
| 
13508
 | 
   141  | 
  | 
| 
23746
 | 
   142  | 
| Propose: "[| evsp:p2; Says A' B {|Agent A,Number r,I,cons M L|}:set evsp;
 | 
| 
 | 
   143  | 
  I:agl; J:agl; isin (Agent C, app (J, del (Agent B, I)));
  | 
| 
 | 
   144  | 
  Nonce ofr ~:used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp : p2"
  | 
| 
13508
 | 
   145  | 
  | 
| 
 | 
   146  | 
subsubsection{*valid offer lists*}
 | 
| 
 | 
   147  | 
  | 
| 
23746
 | 
   148  | 
inductive_set
  | 
| 
 | 
   149  | 
  valid :: "agent => nat => agent => msg set"
  | 
| 
 | 
   150  | 
  for A :: agent and  n :: nat and B :: agent
  | 
| 
 | 
   151  | 
where
  | 
| 
 | 
   152  | 
  Request [intro]: "cons (anchor A n B) nil:valid A n B"
  | 
| 
13508
 | 
   153  | 
  | 
| 
23746
 | 
   154  | 
| Propose [intro]: "L:valid A n B
  | 
| 
 | 
   155  | 
  ==> cons (chain (next_shop (head L)) ofr A L C) L:valid A n B"
  | 
| 
13508
 | 
   156  | 
  | 
| 
 | 
   157  | 
subsubsection{*basic properties of valid*}
 | 
| 
 | 
   158  | 
  | 
| 
 | 
   159  | 
lemma valid_not_empty: "L:valid A n B ==> EX M L'. L = cons M L'"
  | 
| 
 | 
   160  | 
by (erule valid.cases, auto)
  | 
| 
 | 
   161  | 
  | 
| 
 | 
   162  | 
lemma valid_pos_len: "L:valid A n B ==> 0 < len L"
  | 
| 
 | 
   163  | 
by (erule valid.induct, auto)
  | 
| 
 | 
   164  | 
  | 
| 
 | 
   165  | 
subsubsection{*list of offers*}
 | 
| 
 | 
   166  | 
  | 
| 
 | 
   167  | 
consts offers :: "msg => msg"
  | 
| 
 | 
   168  | 
  | 
| 
 | 
   169  | 
recdef offers "measure size"
  | 
| 
 | 
   170  | 
"offers (cons M L) = cons {|shop M, nonce M|} (offers L)"
 | 
| 
 | 
   171  | 
"offers other = nil"
  | 
| 
 | 
   172  | 
  | 
| 
 | 
   173  | 
  | 
| 
 | 
   174  | 
subsection{*Properties of Protocol P2*}
 | 
| 
 | 
   175  | 
  | 
| 
 | 
   176  | 
text{*same as @{text P1_Prop} except that publicly verifiable forward
 | 
| 
 | 
   177  | 
integrity is replaced by forward privacy*}
  | 
| 
 | 
   178  | 
  | 
| 
 | 
   179  | 
subsection{*strong forward integrity:
 | 
| 
 | 
   180  | 
except the last one, no offer can be modified*}
  | 
| 
 | 
   181  | 
  | 
| 
 | 
   182  | 
lemma strong_forward_integrity: "ALL L. Suc i < len L
  | 
| 
 | 
   183  | 
--> L:valid A n B --> repl (L,Suc i,M):valid A n B --> M = ith (L,Suc i)"
  | 
| 
 | 
   184  | 
apply (induct i)
  | 
| 
 | 
   185  | 
(* i = 0 *)
  | 
| 
 | 
   186  | 
apply clarify
  | 
| 
 | 
   187  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
 | 
   188  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   189  | 
apply (ind_cases "{|x,xa,l'a|}:valid A n B" for x xa l'a)
 | 
| 
 | 
   190  | 
apply (ind_cases "{|x,M,l'a|}:valid A n B" for x l'a)
 | 
| 
13508
 | 
   191  | 
apply (simp add: chain_def)
  | 
| 
 | 
   192  | 
(* i > 0 *)
  | 
| 
 | 
   193  | 
apply clarify
  | 
| 
 | 
   194  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   195  | 
apply (ind_cases "{|x,repl(l',Suc na,M)|}:valid A n B" for x l' na)
 | 
| 
13508
 | 
   196  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   197  | 
apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 
13508
 | 
   198  | 
by (drule_tac x=l' in spec, simp, blast)
  | 
| 
 | 
   199  | 
  | 
| 
 | 
   200  | 
subsection{*insertion resilience:
 | 
| 
 | 
   201  | 
except at the beginning, no offer can be inserted*}
  | 
| 
 | 
   202  | 
  | 
| 
 | 
   203  | 
lemma chain_isnt_head [simp]: "L:valid A n B ==>
  | 
| 
 | 
   204  | 
head L ~= chain (next_shop (head L)) ofr A L C"
  | 
| 
 | 
   205  | 
by (erule valid.induct, auto simp: chain_def sign_def anchor_def)
  | 
| 
 | 
   206  | 
  | 
| 
 | 
   207  | 
lemma insertion_resilience: "ALL L. L:valid A n B --> Suc i < len L
  | 
| 
 | 
   208  | 
--> ins (L,Suc i,M) ~:valid A n B"
  | 
| 
 | 
   209  | 
apply (induct i)
  | 
| 
 | 
   210  | 
(* i = 0 *)
  | 
| 
 | 
   211  | 
apply clarify
  | 
| 
 | 
   212  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   213  | 
apply (ind_cases "{|x,l'|}:valid A n B" for x l', simp)
 | 
| 
 | 
   214  | 
apply (ind_cases "{|x,M,l'|}:valid A n B" for x l', clarsimp)
 | 
| 
 | 
   215  | 
apply (ind_cases "{|head l',l'|}:valid A n B" for l', simp, simp)
 | 
| 
13508
 | 
   216  | 
(* i > 0 *)
  | 
| 
 | 
   217  | 
apply clarify
  | 
| 
 | 
   218  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   219  | 
apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 
13508
 | 
   220  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   221  | 
apply (ind_cases "{|x,ins(l',Suc na,M)|}:valid A n B" for x l' na)
 | 
| 
13508
 | 
   222  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
 | 
   223  | 
by (drule_tac x=l' in spec, clarsimp)
  | 
| 
 | 
   224  | 
  | 
| 
 | 
   225  | 
subsection{*truncation resilience:
 | 
| 
 | 
   226  | 
only shop i can truncate at offer i*}
  | 
| 
 | 
   227  | 
  | 
| 
 | 
   228  | 
lemma truncation_resilience: "ALL L. L:valid A n B --> Suc i < len L
  | 
| 
 | 
   229  | 
--> cons M (trunc (L,Suc i)):valid A n B --> shop M = shop (ith (L,i))"
  | 
| 
 | 
   230  | 
apply (induct i)
  | 
| 
 | 
   231  | 
(* i = 0 *)
  | 
| 
 | 
   232  | 
apply clarify
  | 
| 
 | 
   233  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   234  | 
apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 
13508
 | 
   235  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   236  | 
apply (ind_cases "{|M,l'|}:valid A n B" for l')
 | 
| 
13508
 | 
   237  | 
apply (frule len_not_empty, clarsimp, simp)
  | 
| 
 | 
   238  | 
(* i > 0 *)
  | 
| 
 | 
   239  | 
apply clarify
  | 
| 
 | 
   240  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
23746
 | 
   241  | 
apply (ind_cases "{|x,l'|}:valid A n B" for x l')
 | 
| 
13508
 | 
   242  | 
apply (frule len_not_empty, clarsimp)
  | 
| 
 | 
   243  | 
by (drule_tac x=l' in spec, clarsimp)
  | 
| 
 | 
   244  | 
  | 
| 
 | 
   245  | 
subsection{*declarations for tactics*}
 | 
| 
 | 
   246  | 
  | 
| 
 | 
   247  | 
declare knows_Spy_partsEs [elim]
  | 
| 
 | 
   248  | 
declare Fake_parts_insert [THEN subsetD, dest]
  | 
| 
 | 
   249  | 
declare initState.simps [simp del]
  | 
| 
 | 
   250  | 
  | 
| 
 | 
   251  | 
subsection{*get components of a message*}
 | 
| 
 | 
   252  | 
  | 
| 
 | 
   253  | 
lemma get_ML [dest]: "Says A' B {|A,R,I,M,L|}:set evs ==>
 | 
| 
 | 
   254  | 
M:parts (spies evs) & L:parts (spies evs)"
  | 
| 
 | 
   255  | 
by blast
  | 
| 
 | 
   256  | 
  | 
| 
 | 
   257  | 
subsection{*general properties of p2*}
 | 
| 
 | 
   258  | 
  | 
| 
 | 
   259  | 
lemma reqm_neq_prom [iff]:
  | 
| 
 | 
   260  | 
"reqm A r n I B ~= prom B' ofr A' r' I' (cons M L) J C"
  | 
| 
 | 
   261  | 
by (auto simp: reqm_def prom_def)
  | 
| 
 | 
   262  | 
  | 
| 
 | 
   263  | 
lemma prom_neq_reqm [iff]:
  | 
| 
 | 
   264  | 
"prom B' ofr A' r' I' (cons M L) J C ~= reqm A r n I B"
  | 
| 
 | 
   265  | 
by (auto simp: reqm_def prom_def)
  | 
| 
 | 
   266  | 
  | 
| 
 | 
   267  | 
lemma req_neq_pro [iff]: "req A r n I B ~= pro B' ofr A' r' I' (cons M L) J C"
  | 
| 
 | 
   268  | 
by (auto simp: req_def pro_def)
  | 
| 
 | 
   269  | 
  | 
| 
 | 
   270  | 
lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C ~= req A r n I B"
  | 
| 
 | 
   271  | 
by (auto simp: req_def pro_def)
  | 
| 
 | 
   272  | 
  | 
| 
 | 
   273  | 
lemma p2_has_no_Gets: "evs:p2 ==> ALL A X. Gets A X ~:set evs"
  | 
| 
 | 
   274  | 
by (erule p2.induct, auto simp: req_def pro_def)
  | 
| 
 | 
   275  | 
  | 
| 
 | 
   276  | 
lemma p2_is_Gets_correct [iff]: "Gets_correct p2"
  | 
| 
 | 
   277  | 
by (auto simp: Gets_correct_def dest: p2_has_no_Gets)
  | 
| 
 | 
   278  | 
  | 
| 
 | 
   279  | 
lemma p2_is_one_step [iff]: "one_step p2"
  | 
| 
23746
 | 
   280  | 
by (unfold one_step_def, clarify, ind_cases "ev#evs:p2" for ev evs, auto)
  | 
| 
13508
 | 
   281  | 
  | 
| 
 | 
   282  | 
lemma p2_has_only_Says' [rule_format]: "evs:p2 ==>
  | 
| 
 | 
   283  | 
ev:set evs --> (EX A B X. ev=Says A B X)"
  | 
| 
 | 
   284  | 
by (erule p2.induct, auto simp: req_def pro_def)
  | 
| 
 | 
   285  | 
  | 
| 
 | 
   286  | 
lemma p2_has_only_Says [iff]: "has_only_Says p2"
  | 
| 
 | 
   287  | 
by (auto simp: has_only_Says_def dest: p2_has_only_Says')
  | 
| 
 | 
   288  | 
  | 
| 
 | 
   289  | 
lemma p2_is_regular [iff]: "regular p2"
  | 
| 
 | 
   290  | 
apply (simp only: regular_def, clarify)
  | 
| 
 | 
   291  | 
apply (erule_tac p2.induct)
  | 
| 
 | 
   292  | 
apply (simp_all add: initState.simps knows.simps pro_def prom_def
  | 
| 
 | 
   293  | 
req_def reqm_def anchor_def chain_def sign_def)
  | 
| 
 | 
   294  | 
by (auto dest: no_Key_in_agl no_Key_in_appdel parts_trans)
  | 
| 
 | 
   295  | 
  | 
| 
 | 
   296  | 
subsection{*private keys are safe*}
 | 
| 
 | 
   297  | 
  | 
| 
 | 
   298  | 
lemma priK_parts_Friend_imp_bad [rule_format,dest]:
  | 
| 
 | 
   299  | 
     "[| evs:p2; Friend B ~= A |]
  | 
| 
 | 
   300  | 
      ==> (Key (priK A):parts (knows (Friend B) evs)) --> (A:bad)"
  | 
| 
 | 
   301  | 
apply (erule p2.induct)
  | 
| 
 | 
   302  | 
apply (simp_all add: initState.simps knows.simps pro_def prom_def
  | 
| 
17778
 | 
   303  | 
                req_def reqm_def anchor_def chain_def sign_def) 
  | 
| 
13508
 | 
   304  | 
apply (blast dest: no_Key_in_agl)
  | 
| 
 | 
   305  | 
apply (auto del: parts_invKey disjE  dest: parts_trans
  | 
| 
 | 
   306  | 
            simp add: no_Key_in_appdel)
  | 
| 
 | 
   307  | 
done
  | 
| 
 | 
   308  | 
  | 
| 
 | 
   309  | 
lemma priK_analz_Friend_imp_bad [rule_format,dest]:
  | 
| 
 | 
   310  | 
     "[| evs:p2; Friend B ~= A |]
  | 
| 
 | 
   311  | 
==> (Key (priK A):analz (knows (Friend B) evs)) --> (A:bad)"
  | 
| 
 | 
   312  | 
by auto
  | 
| 
 | 
   313  | 
  | 
| 
 | 
   314  | 
lemma priK_notin_knows_max_Friend:
  | 
| 
 | 
   315  | 
     "[| evs:p2; A ~:bad; A ~= Friend C |]
  | 
| 
 | 
   316  | 
      ==> Key (priK A) ~:analz (knows_max (Friend C) evs)"
  | 
| 
 | 
   317  | 
apply (rule not_parts_not_analz, simp add: knows_max_def, safe)
  | 
| 
 | 
   318  | 
apply (drule_tac H="spies' evs" in parts_sub)
  | 
| 
 | 
   319  | 
apply (rule_tac p=p2 in knows_max'_sub_spies', simp+)
  | 
| 
 | 
   320  | 
apply (drule_tac H="spies evs" in parts_sub)
  | 
| 
 | 
   321  | 
by (auto dest: knows'_sub_knows [THEN subsetD] priK_notin_initState_Friend)
  | 
| 
 | 
   322  | 
  | 
| 
 | 
   323  | 
subsection{*general guardedness properties*}
 | 
| 
 | 
   324  | 
  | 
| 
 | 
   325  | 
lemma agl_guard [intro]: "I:agl ==> I:guard n Ks"
  | 
| 
 | 
   326  | 
by (erule agl.induct, auto)
  | 
| 
 | 
   327  | 
  | 
| 
 | 
   328  | 
lemma Says_to_knows_max'_guard: "[| Says A' C {|A'',r,I,L|}:set evs;
 | 
| 
 | 
   329  | 
Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks"
  | 
| 
 | 
   330  | 
by (auto dest: Says_to_knows_max')
  | 
| 
 | 
   331  | 
  | 
| 
 | 
   332  | 
lemma Says_from_knows_max'_guard: "[| Says C A' {|A'',r,I,L|}:set evs;
 | 
| 
 | 
   333  | 
Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks"
  | 
| 
 | 
   334  | 
by (auto dest: Says_from_knows_max')
  | 
| 
 | 
   335  | 
  | 
| 
 | 
   336  | 
lemma Says_Nonce_not_used_guard: "[| Says A' B {|A'',r,I,L|}:set evs;
 | 
| 
 | 
   337  | 
Nonce n ~:used evs |] ==> L:guard n Ks"
  | 
| 
 | 
   338  | 
by (drule not_used_not_parts, auto)
  | 
| 
 | 
   339  | 
  | 
| 
 | 
   340  | 
subsection{*guardedness of messages*}
 | 
| 
 | 
   341  | 
  | 
| 
 | 
   342  | 
lemma chain_guard [iff]: "chain B ofr A L C:guard n {priK A}"
 | 
| 
 | 
   343  | 
by (case_tac "ofr=n", auto simp: chain_def sign_def)
  | 
| 
 | 
   344  | 
  | 
| 
 | 
   345  | 
lemma chain_guard_Nonce_neq [intro]: "n ~= ofr
  | 
| 
 | 
   346  | 
==> chain B ofr A' L C:guard n {priK A}"
 | 
| 
 | 
   347  | 
by (auto simp: chain_def sign_def)
  | 
| 
 | 
   348  | 
  | 
| 
 | 
   349  | 
lemma anchor_guard [iff]: "anchor A n' B:guard n {priK A}"
 | 
| 
 | 
   350  | 
by (case_tac "n'=n", auto simp: anchor_def)
  | 
| 
 | 
   351  | 
  | 
| 
 | 
   352  | 
lemma anchor_guard_Nonce_neq [intro]: "n ~= n'
  | 
| 
 | 
   353  | 
==> anchor A' n' B:guard n {priK A}"
 | 
| 
 | 
   354  | 
by (auto simp: anchor_def)
  | 
| 
 | 
   355  | 
  | 
| 
 | 
   356  | 
lemma reqm_guard [intro]: "I:agl ==> reqm A r n' I B:guard n {priK A}"
 | 
| 
 | 
   357  | 
by (case_tac "n'=n", auto simp: reqm_def)
  | 
| 
 | 
   358  | 
  | 
| 
 | 
   359  | 
lemma reqm_guard_Nonce_neq [intro]: "[| n ~= n'; I:agl |]
  | 
| 
 | 
   360  | 
==> reqm A' r n' I B:guard n {priK A}"
 | 
| 
 | 
   361  | 
by (auto simp: reqm_def)
  | 
| 
 | 
   362  | 
  | 
| 
 | 
   363  | 
lemma prom_guard [intro]: "[| I:agl; J:agl; L:guard n {priK A} |]
 | 
| 
 | 
   364  | 
==> prom B ofr A r I L J C:guard n {priK A}"
 | 
| 
 | 
   365  | 
by (auto simp: prom_def)
  | 
| 
 | 
   366  | 
  | 
| 
 | 
   367  | 
lemma prom_guard_Nonce_neq [intro]: "[| n ~= ofr; I:agl; J:agl;
  | 
| 
 | 
   368  | 
L:guard n {priK A} |] ==> prom B ofr A' r I L J C:guard n {priK A}"
 | 
| 
 | 
   369  | 
by (auto simp: prom_def)
  | 
| 
 | 
   370  | 
  | 
| 
 | 
   371  | 
subsection{*Nonce uniqueness*}
 | 
| 
 | 
   372  | 
  | 
| 
 | 
   373  | 
lemma uniq_Nonce_in_chain [dest]: "Nonce k:parts {chain B ofr A L C} ==> k=ofr"
 | 
| 
 | 
   374  | 
by (auto simp: chain_def sign_def)
  | 
| 
 | 
   375  | 
  | 
| 
 | 
   376  | 
lemma uniq_Nonce_in_anchor [dest]: "Nonce k:parts {anchor A n B} ==> k=n"
 | 
| 
 | 
   377  | 
by (auto simp: anchor_def chain_def sign_def)
  | 
| 
 | 
   378  | 
  | 
| 
 | 
   379  | 
lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k:parts {reqm A r n I B};
 | 
| 
 | 
   380  | 
I:agl |] ==> k=n"
  | 
| 
 | 
   381  | 
by (auto simp: reqm_def dest: no_Nonce_in_agl)
  | 
| 
 | 
   382  | 
  | 
| 
 | 
   383  | 
lemma uniq_Nonce_in_prom [dest]: "[| Nonce k:parts {prom B ofr A r I L J C};
 | 
| 
 | 
   384  | 
I:agl; J:agl; Nonce k ~:parts {L} |] ==> k=ofr"
 | 
| 
 | 
   385  | 
by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel)
  | 
| 
 | 
   386  | 
  | 
| 
 | 
   387  | 
subsection{*requests are guarded*}
 | 
| 
 | 
   388  | 
  | 
| 
 | 
   389  | 
lemma req_imp_Guard [rule_format]: "[| evs:p2; A ~:bad |] ==>
  | 
| 
 | 
   390  | 
req A r n I B:set evs --> Guard n {priK A} (spies evs)"
 | 
| 
 | 
   391  | 
apply (erule p2.induct, simp)
  | 
| 
 | 
   392  | 
apply (simp add: req_def knows.simps, safe)
  | 
| 
 | 
   393  | 
apply (erule in_synth_Guard, erule Guard_analz, simp)
  | 
| 
 | 
   394  | 
by (auto simp: req_def pro_def dest: Says_imp_knows_Spy)
  | 
| 
 | 
   395  | 
  | 
| 
 | 
   396  | 
lemma req_imp_Guard_Friend: "[| evs:p2; A ~:bad; req A r n I B:set evs |]
  | 
| 
 | 
   397  | 
==> Guard n {priK A} (knows_max (Friend C) evs)"
 | 
| 
 | 
   398  | 
apply (rule Guard_knows_max')
  | 
| 
 | 
   399  | 
apply (rule_tac H="spies evs" in Guard_mono)
  | 
| 
 | 
   400  | 
apply (rule req_imp_Guard, simp+)
  | 
| 
 | 
   401  | 
apply (rule_tac B="spies' evs" in subset_trans)
  | 
| 
 | 
   402  | 
apply (rule_tac p=p2 in knows_max'_sub_spies', simp+)
  | 
| 
 | 
   403  | 
by (rule knows'_sub_knows)
  | 
| 
 | 
   404  | 
  | 
| 
 | 
   405  | 
subsection{*propositions are guarded*}
 | 
| 
 | 
   406  | 
  | 
| 
 | 
   407  | 
lemma pro_imp_Guard [rule_format]: "[| evs:p2; B ~:bad; A ~:bad |] ==>
  | 
| 
 | 
   408  | 
pro B ofr A r I (cons M L) J C:set evs --> Guard ofr {priK A} (spies evs)"
 | 
| 
 | 
   409  | 
apply (erule p2.induct) (* +3 subgoals *)
  | 
| 
 | 
   410  | 
(* Nil *)
  | 
| 
 | 
   411  | 
apply simp
  | 
| 
 | 
   412  | 
(* Fake *)
  | 
| 
 | 
   413  | 
apply (simp add: pro_def, safe) (* +4 subgoals *)
  | 
| 
 | 
   414  | 
(* 1 *)
  | 
| 
 | 
   415  | 
apply (erule in_synth_Guard, drule Guard_analz, simp, simp)
  | 
| 
 | 
   416  | 
(* 2 *)
  | 
| 
 | 
   417  | 
apply simp
  | 
| 
 | 
   418  | 
(* 3 *)
  | 
| 
 | 
   419  | 
apply (simp, simp add: req_def pro_def, blast)
  | 
| 
 | 
   420  | 
(* 4 *)
  | 
| 
 | 
   421  | 
apply (simp add: pro_def)
  | 
| 
 | 
   422  | 
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard)
  | 
| 
 | 
   423  | 
(* 5 *)
  | 
| 
 | 
   424  | 
apply simp
  | 
| 
 | 
   425  | 
apply safe (* +1 subgoal *)
  | 
| 
 | 
   426  | 
apply (simp add: pro_def)
  | 
| 
 | 
   427  | 
apply (blast dest: prom_inj Says_Nonce_not_used_guard)
  | 
| 
 | 
   428  | 
(* 6 *)
  | 
| 
 | 
   429  | 
apply (simp add: pro_def)
  | 
| 
 | 
   430  | 
apply (blast dest: Says_imp_knows_Spy)
  | 
| 
 | 
   431  | 
(* Request *)
  | 
| 
 | 
   432  | 
apply (simp add: pro_def)
  | 
| 
 | 
   433  | 
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard)
  | 
| 
 | 
   434  | 
(* Propose *)
  | 
| 
 | 
   435  | 
apply simp
  | 
| 
 | 
   436  | 
apply safe (* +1 subgoal *)
  | 
| 
 | 
   437  | 
(* 1 *)
  | 
| 
 | 
   438  | 
apply (simp add: pro_def)
  | 
| 
 | 
   439  | 
apply (blast dest: prom_inj Says_Nonce_not_used_guard)
  | 
| 
 | 
   440  | 
(* 2 *)
  | 
| 
 | 
   441  | 
apply (simp add: pro_def)
  | 
| 
 | 
   442  | 
by (blast dest: Says_imp_knows_Spy)
  | 
| 
 | 
   443  | 
  | 
| 
 | 
   444  | 
lemma pro_imp_Guard_Friend: "[| evs:p2; B ~:bad; A ~:bad;
  | 
| 
 | 
   445  | 
pro B ofr A r I (cons M L) J C:set evs |]
  | 
| 
 | 
   446  | 
==> Guard ofr {priK A} (knows_max (Friend D) evs)"
 | 
| 
 | 
   447  | 
apply (rule Guard_knows_max')
  | 
| 
 | 
   448  | 
apply (rule_tac H="spies evs" in Guard_mono)
  | 
| 
 | 
   449  | 
apply (rule pro_imp_Guard, simp+)
  | 
| 
 | 
   450  | 
apply (rule_tac B="spies' evs" in subset_trans)
  | 
| 
 | 
   451  | 
apply (rule_tac p=p2 in knows_max'_sub_spies', simp+)
  | 
| 
 | 
   452  | 
by (rule knows'_sub_knows)
  | 
| 
 | 
   453  | 
  | 
| 
 | 
   454  | 
subsection{*data confidentiality:
 | 
| 
 | 
   455  | 
no one other than the originator can decrypt the offers*}
  | 
| 
 | 
   456  | 
  | 
| 
 | 
   457  | 
lemma Nonce_req_notin_spies: "[| evs:p2; req A r n I B:set evs; A ~:bad |]
  | 
| 
 | 
   458  | 
==> Nonce n ~:analz (spies evs)"
  | 
| 
 | 
   459  | 
by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
  | 
| 
 | 
   460  | 
  | 
| 
 | 
   461  | 
lemma Nonce_req_notin_knows_max_Friend: "[| evs:p2; req A r n I B:set evs;
  | 
| 
 | 
   462  | 
A ~:bad; A ~= Friend C |] ==> Nonce n ~:analz (knows_max (Friend C) evs)"
  | 
| 
 | 
   463  | 
apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+)
  | 
| 
 | 
   464  | 
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+)
  | 
| 
 | 
   465  | 
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def)
  | 
| 
 | 
   466  | 
  | 
| 
 | 
   467  | 
lemma Nonce_pro_notin_spies: "[| evs:p2; B ~:bad; A ~:bad;
  | 
| 
 | 
   468  | 
pro B ofr A r I (cons M L) J C:set evs |] ==> Nonce ofr ~:analz (spies evs)"
  | 
| 
 | 
   469  | 
by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+)
  | 
| 
 | 
   470  | 
  | 
| 
 | 
   471  | 
lemma Nonce_pro_notin_knows_max_Friend: "[| evs:p2; B ~:bad; A ~:bad;
  | 
| 
 | 
   472  | 
A ~= Friend D; pro B ofr A r I (cons M L) J C:set evs |]
  | 
| 
 | 
   473  | 
==> Nonce ofr ~:analz (knows_max (Friend D) evs)"
  | 
| 
 | 
   474  | 
apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+)
  | 
| 
 | 
   475  | 
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+)
  | 
| 
 | 
   476  | 
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def)
  | 
| 
 | 
   477  | 
  | 
| 
 | 
   478  | 
subsection{*forward privacy:
 | 
| 
 | 
   479  | 
only the originator can know the identity of the shops*}
  | 
| 
 | 
   480  | 
  | 
| 
 | 
   481  | 
lemma forward_privacy_Spy: "[| evs:p2; B ~:bad; A ~:bad;
  | 
| 
 | 
   482  | 
pro B ofr A r I (cons M L) J C:set evs |]
  | 
| 
 | 
   483  | 
==> sign B (Nonce ofr) ~:analz (spies evs)"
  | 
| 
 | 
   484  | 
by (auto simp:sign_def dest: Nonce_pro_notin_spies)
  | 
| 
 | 
   485  | 
  | 
| 
 | 
   486  | 
lemma forward_privacy_Friend: "[| evs:p2; B ~:bad; A ~:bad; A ~= Friend D;
  | 
| 
 | 
   487  | 
pro B ofr A r I (cons M L) J C:set evs |]
  | 
| 
 | 
   488  | 
==> sign B (Nonce ofr) ~:analz (knows_max (Friend D) evs)"
  | 
| 
 | 
   489  | 
by (auto simp:sign_def dest:Nonce_pro_notin_knows_max_Friend )
  | 
| 
 | 
   490  | 
  | 
| 
 | 
   491  | 
subsection{*non repudiability: an offer signed by B has been sent by B*}
 | 
| 
 | 
   492  | 
  | 
| 
 | 
   493  | 
lemma Crypt_reqm: "[| Crypt (priK A) X:parts {reqm A' r n I B}; I:agl |] ==> A=A'"
 | 
| 
 | 
   494  | 
by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl)
  | 
| 
 | 
   495  | 
  | 
| 
 | 
   496  | 
lemma Crypt_prom: "[| Crypt (priK A) X:parts {prom B ofr A' r I L J C};
 | 
| 
 | 
   497  | 
I:agl; J:agl |] ==> A=B | Crypt (priK A) X:parts {L}"
 | 
| 
 | 
   498  | 
apply (simp add: prom_def anchor_def chain_def sign_def)
  | 
| 
 | 
   499  | 
by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel)
  | 
| 
 | 
   500  | 
  | 
| 
 | 
   501  | 
lemma Crypt_safeness: "[| evs:p2; A ~:bad |] ==> Crypt (priK A) X:parts (spies evs)
  | 
| 
 | 
   502  | 
--> (EX B Y. Says A B Y:set evs & Crypt (priK A) X:parts {Y})"
 | 
| 
 | 
   503  | 
apply (erule p2.induct)
  | 
| 
 | 
   504  | 
(* Nil *)
  | 
| 
 | 
   505  | 
apply simp
  | 
| 
 | 
   506  | 
(* Fake *)
  | 
| 
 | 
   507  | 
apply clarsimp
  | 
| 
 | 
   508  | 
apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp)
  | 
| 
 | 
   509  | 
apply (erule disjE)
  | 
| 
 | 
   510  | 
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast)
  | 
| 
 | 
   511  | 
(* Request *)
  | 
| 
 | 
   512  | 
apply (simp add: req_def, clarify)
  | 
| 
 | 
   513  | 
apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp)
  | 
| 
 | 
   514  | 
apply (erule disjE)
  | 
| 
 | 
   515  | 
apply (frule Crypt_reqm, simp, clarify)
  | 
| 
 | 
   516  | 
apply (rule_tac x=B in exI, rule_tac x="reqm A r n I B" in exI, simp, blast)
  | 
| 
 | 
   517  | 
(* Propose *)
  | 
| 
 | 
   518  | 
apply (simp add: pro_def, clarify)
  | 
| 
 | 
   519  | 
apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp)
  | 
| 
 | 
   520  | 
apply (rotate_tac -1, erule disjE)
  | 
| 
 | 
   521  | 
apply (frule Crypt_prom, simp, simp)
  | 
| 
 | 
   522  | 
apply (rotate_tac -1, erule disjE)
  | 
| 
 | 
   523  | 
apply (rule_tac x=C in exI)
  | 
| 
 | 
   524  | 
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI, blast)
  | 
| 
 | 
   525  | 
apply (subgoal_tac "cons M L:parts (spies evsp)")
  | 
| 
 | 
   526  | 
apply (drule_tac G="{cons M L}" and H="spies evsp" in parts_trans, blast, blast)
 | 
| 
 | 
   527  | 
apply (drule Says_imp_spies, rotate_tac -1, drule parts.Inj)
  | 
| 
 | 
   528  | 
apply (drule parts.Snd, drule parts.Snd, drule parts.Snd)
  | 
| 
 | 
   529  | 
by auto
  | 
| 
 | 
   530  | 
  | 
| 
 | 
   531  | 
lemma Crypt_Hash_imp_sign: "[| evs:p2; A ~:bad |] ==>
  | 
| 
 | 
   532  | 
Crypt (priK A) (Hash X):parts (spies evs)
  | 
| 
 | 
   533  | 
--> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})"
 | 
| 
 | 
   534  | 
apply (erule p2.induct)
  | 
| 
 | 
   535  | 
(* Nil *)
  | 
| 
 | 
   536  | 
apply simp
  | 
| 
 | 
   537  | 
(* Fake *)
  | 
| 
 | 
   538  | 
apply clarsimp
  | 
| 
 | 
   539  | 
apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD)
  | 
| 
 | 
   540  | 
apply simp
  | 
| 
 | 
   541  | 
apply (erule disjE)
  | 
| 
 | 
   542  | 
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast)
  | 
| 
 | 
   543  | 
(* Request *)
  | 
| 
 | 
   544  | 
apply (simp add: req_def, clarify)
  | 
| 
 | 
   545  | 
apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD)
  | 
| 
 | 
   546  | 
apply simp
  | 
| 
 | 
   547  | 
apply (erule disjE)
  | 
| 
 | 
   548  | 
apply (frule Crypt_reqm, simp+)
  | 
| 
 | 
   549  | 
apply (rule_tac x=B in exI, rule_tac x="reqm Aa r n I B" in exI)
  | 
| 
 | 
   550  | 
apply (simp add: reqm_def sign_def anchor_def no_Crypt_in_agl)
  | 
| 
 | 
   551  | 
apply (simp add: chain_def sign_def, blast)
  | 
| 
 | 
   552  | 
(* Propose *)
  | 
| 
 | 
   553  | 
apply (simp add: pro_def, clarify)
  | 
| 
 | 
   554  | 
apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD)
  | 
| 
 | 
   555  | 
apply simp
  | 
| 
 | 
   556  | 
apply (rotate_tac -1, erule disjE)
  | 
| 
 | 
   557  | 
apply (simp add: prom_def sign_def no_Crypt_in_agl no_Crypt_in_appdel)
  | 
| 
 | 
   558  | 
apply (simp add: chain_def sign_def)
  | 
| 
 | 
   559  | 
apply (rotate_tac -1, erule disjE)
  | 
| 
 | 
   560  | 
apply (rule_tac x=C in exI)
  | 
| 
 | 
   561  | 
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI)
  | 
| 
 | 
   562  | 
apply (simp add: prom_def chain_def sign_def)
  | 
| 
 | 
   563  | 
apply (erule impE) 
  | 
| 
 | 
   564  | 
apply (blast dest: get_ML parts_sub) 
  | 
| 
 | 
   565  | 
apply (blast del: MPair_parts)+
  | 
| 
 | 
   566  | 
done
  | 
| 
 | 
   567  | 
  | 
| 
 | 
   568  | 
lemma sign_safeness: "[| evs:p2; A ~:bad |] ==> sign A X:parts (spies evs)
  | 
| 
 | 
   569  | 
--> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})"
 | 
| 
 | 
   570  | 
apply (clarify, simp add: sign_def, frule parts.Snd)
  | 
| 
 | 
   571  | 
apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def])
  | 
| 
 | 
   572  | 
done
  | 
| 
 | 
   573  | 
  | 
| 
 | 
   574  | 
end  |