src/HOL/SET-Protocol/Cardholder_Registration.thy
author paulson
Fri, 26 Sep 2003 10:32:26 +0200
changeset 14206 77bf175f5145
parent 14199 d3b8d972a488
child 14218 db95d1c2f51b
permissions -rw-r--r--
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/SET/Cardholder_Registration
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     2
    ID:         $Id$
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     3
    Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     4
                Piero Tramontano
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     5
*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     6
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     7
header{*The SET Cardholder Registration Protocol*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     8
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     9
theory Cardholder_Registration = PublicSET:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    10
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    11
text{*Note: nonces seem to consist of 20 bytes.  That includes both freshness
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    12
challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    13
*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    14
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    15
text{*Simplifications involving @{text analz_image_keys_simps} appear to
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    16
have become much slower. The cause is unclear. However, there is a big blow-up
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    17
and the rewriting is very sensitive to the set of rewrite rules given.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    18
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    19
subsection{*Predicate Formalizing the Encryption Association between Keys *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    20
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    21
consts
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    22
  KeyCryptKey :: "[key, key, event list] => bool"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    23
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    24
primrec
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    25
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    26
KeyCryptKey_Nil:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    27
  "KeyCryptKey DK K [] = False"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    28
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    29
KeyCryptKey_Cons:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    30
      --{*Says is the only important case.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    31
	1st case: CR5, where KC3 encrypts KC2.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    32
	2nd case: any use of priEK C.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    33
	Revision 1.12 has a more complicated version with separate treatment of
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    34
	  the dependency of KC1, KC2 and KC3 on priEK (CA i.)  Not needed since
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    35
	  priEK C is never sent (and so can't be lost except at the start). *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    36
  "KeyCryptKey DK K (ev # evs) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    37
   (KeyCryptKey DK K evs |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    38
    (case ev of
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    39
      Says A B Z =>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    40
       ((\<exists>N X Y. A \<noteq> Spy &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    41
	         DK \<in> symKeys &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    42
		 Z = {|Crypt DK {|Agent A, Nonce N, Key K, X|}, Y|}) |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    43
	(\<exists>C. DK = priEK C))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    44
    | Gets A' X => False
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    45
    | Notes A' X => False))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    46
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    47
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    48
subsection{*Predicate formalizing the association between keys and nonces *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    49
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    50
consts
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    51
  KeyCryptNonce :: "[key, key, event list] => bool"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    52
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    53
primrec
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    54
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    55
KeyCryptNonce_Nil:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    56
  "KeyCryptNonce EK K [] = False"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    57
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    58
KeyCryptNonce_Cons:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    59
  --{*Says is the only important case.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    60
    1st case: CR3, where KC1 encrypts NC2 (distinct from CR5 due to EXH);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    61
    2nd case: CR5, where KC3 encrypts NC3;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    62
    3rd case: CR6, where KC2 encrypts NC3;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    63
    4th case: CR6, where KC2 encrypts NonceCCA;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    64
    5th case: any use of @{term "priEK C"} (including CardSecret).
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    65
    NB the only Nonces we need to keep secret are CardSecret and NonceCCA.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    66
    But we can't prove @{text Nonce_compromise} unless the relation covers ALL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    67
	nonces that the protocol keeps secret.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    68
  *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    69
  "KeyCryptNonce DK N (ev # evs) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    70
   (KeyCryptNonce DK N evs |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    71
    (case ev of
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    72
      Says A B Z =>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    73
       A \<noteq> Spy &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    74
       ((\<exists>X Y. DK \<in> symKeys &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    75
	       Z = (EXHcrypt DK X {|Agent A, Nonce N|} Y)) |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    76
	(\<exists>X Y. DK \<in> symKeys &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    77
	       Z = {|Crypt DK {|Agent A, Nonce N, X|}, Y|}) |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    78
	(\<exists>K i X Y.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    79
	  K \<in> symKeys &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    80
          Z = Crypt K {|sign (priSK (CA i)) {|Agent B, Nonce N, X|}, Y|} &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    81
	  (DK=K | KeyCryptKey DK K evs)) |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    82
	(\<exists>K C NC3 Y.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    83
	  K \<in> symKeys &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    84
          Z = Crypt K
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    85
 	        {|sign (priSK C) {|Agent B, Nonce NC3, Agent C, Nonce N|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    86
                  Y|} &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    87
	  (DK=K | KeyCryptKey DK K evs)) |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    88
	(\<exists>C. DK = priEK C))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    89
    | Gets A' X => False
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    90
    | Notes A' X => False))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    91
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    92
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    93
subsection{*Formal protocol definition *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    94
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    95
consts  set_cr  :: "event list set"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    96
inductive set_cr
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    97
 intros
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    98
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    99
  Nil:    --{*Initial trace is empty*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   100
	  "[] \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   101
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   102
  Fake:    --{*The spy MAY say anything he CAN say.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   103
	   "[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   104
	    ==> Says Spy B X  # evsf \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   105
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   106
  Reception: --{*If A sends a message X to B, then B might receive it*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   107
	     "[| evsr \<in> set_cr; Says A B X \<in> set evsr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   108
              ==> Gets B X  # evsr \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   109
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   110
  SET_CR1: --{*CardCInitReq: C initiates a run, sending a nonce to CCA*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   111
	     "[| evs1 \<in> set_cr;  C = Cardholder k;  Nonce NC1 \<notin> used evs1 |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   112
	      ==> Says C (CA i) {|Agent C, Nonce NC1|} # evs1 \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   113
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   114
  SET_CR2: --{*CardCInitRes: CA responds sending NC1 and its certificates*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   115
	     "[| evs2 \<in> set_cr;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   116
		 Gets (CA i) {|Agent C, Nonce NC1|} \<in> set evs2 |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   117
	      ==> Says (CA i) C
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   118
		       {|sign (priSK (CA i)) {|Agent C, Nonce NC1|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   119
			 cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   120
			 cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   121
		    # evs2 \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   122
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   123
  SET_CR3:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   124
   --{*RegFormReq: C sends his PAN and a new nonce to CA.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   125
   C verifies that
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   126
    - nonce received is the same as that sent;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   127
    - certificates are signed by RCA;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   128
    - certificates are an encryption certificate (flag is onlyEnc) and a
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   129
      signature certificate (flag is onlySig);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   130
    - certificates pertain to the CA that C contacted (this is done by
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   131
      checking the signature).
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   132
   C generates a fresh symmetric key KC1.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   133
   The point of encrypting @{term "{|Agent C, Nonce NC2, Hash (Pan(pan C))|}"}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   134
   is not clear. *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   135
"[| evs3 \<in> set_cr;  C = Cardholder k;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   136
    Nonce NC2 \<notin> used evs3;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   137
    Key KC1 \<notin> used evs3; KC1 \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   138
    Gets C {|sign (invKey SKi) {|Agent X, Nonce NC1|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   139
	     cert (CA i) EKi onlyEnc (priSK RCA),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   140
	     cert (CA i) SKi onlySig (priSK RCA)|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   141
       \<in> set evs3;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   142
    Says C (CA i) {|Agent C, Nonce NC1|} \<in> set evs3|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   143
 ==> Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   144
       # Notes C {|Key KC1, Agent (CA i)|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   145
       # evs3 \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   146
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   147
  SET_CR4:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   148
    --{*RegFormRes:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   149
    CA responds sending NC2 back with a new nonce NCA, after checking that
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   150
     - the digital envelope is correctly encrypted by @{term "pubEK (CA i)"}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   151
     - the entire message is encrypted with the same key found inside the
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   152
       envelope (here, KC1) *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   153
"[| evs4 \<in> set_cr;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   154
    Nonce NCA \<notin> used evs4;  KC1 \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   155
    Gets (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan X)))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   156
       \<in> set evs4 |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   157
  ==> Says (CA i) C
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   158
	  {|sign (priSK (CA i)) {|Agent C, Nonce NC2, Nonce NCA|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   159
	    cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   160
	    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   161
       # evs4 \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   162
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   163
  SET_CR5:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   164
   --{*CertReq: C sends his PAN, a new nonce, its proposed public signature key
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   165
       and its half of the secret value to CA.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   166
       We now assume that C has a fixed key pair, and he submits (pubSK C).
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   167
       The protocol does not require this key to be fresh.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   168
       The encryption below is actually EncX.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   169
"[| evs5 \<in> set_cr;  C = Cardholder k;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   170
    Nonce NC3 \<notin> used evs5;  Nonce CardSecret \<notin> used evs5; NC3\<noteq>CardSecret;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   171
    Key KC2 \<notin> used evs5; KC2 \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   172
    Key KC3 \<notin> used evs5; KC3 \<in> symKeys; KC2\<noteq>KC3;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   173
    Gets C {|sign (invKey SKi) {|Agent C, Nonce NC2, Nonce NCA|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   174
             cert (CA i) EKi onlyEnc (priSK RCA),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   175
             cert (CA i) SKi onlySig (priSK RCA) |}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   176
        \<in> set evs5;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   177
    Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   178
         \<in> set evs5 |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   179
==> Says C (CA i)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   180
         {|Crypt KC3
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   181
	     {|Agent C, Nonce NC3, Key KC2, Key (pubSK C),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   182
	       Crypt (priSK C)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   183
	         (Hash {|Agent C, Nonce NC3, Key KC2,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   184
			 Key (pubSK C), Pan (pan C), Nonce CardSecret|})|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   185
           Crypt EKi {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   186
    # Notes C {|Key KC2, Agent (CA i)|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   187
    # Notes C {|Key KC3, Agent (CA i)|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   188
    # evs5 \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   189
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   190
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   191
  --{* CertRes: CA responds sending NC3 back with its half of the secret value,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   192
   its signature certificate and the new cardholder signature
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   193
   certificate.  CA checks to have never certified the key proposed by C.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   194
   NOTE: In Merchant Registration, the corresponding rule (4)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   195
   uses the "sign" primitive. The encryption below is actually @{term EncK}, 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   196
   which is just @{term "Crypt K (sign SK X)"}.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   197
*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   198
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   199
SET_CR6:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   200
"[| evs6 \<in> set_cr;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   201
    Nonce NonceCCA \<notin> used evs6;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   202
    KC2 \<in> symKeys;  KC3 \<in> symKeys;  cardSK \<notin> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   203
    Notes (CA i) (Key cardSK) \<notin> set evs6;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   204
    Gets (CA i)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   205
      {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, Key cardSK,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   206
                    Crypt (invKey cardSK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   207
                      (Hash {|Agent C, Nonce NC3, Key KC2,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   208
			      Key cardSK, Pan (pan C), Nonce CardSecret|})|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   209
        Crypt (pubEK (CA i)) {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   210
      \<in> set evs6 |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   211
==> Says (CA i) C
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   212
         (Crypt KC2
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   213
	  {|sign (priSK (CA i))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   214
	         {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   215
	    certC (pan C) cardSK (XOR(CardSecret,NonceCCA)) onlySig (priSK (CA i)),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   216
	    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   217
      # Notes (CA i) (Key cardSK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   218
      # evs6 \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   219
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   220
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   221
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   222
declare parts.Body [dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   223
declare analz_into_parts [dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   224
declare Fake_parts_insert_in_Un [dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   225
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   226
text{*A "possibility property": there are traces that reach the end.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   227
      An unconstrained proof with many subgoals.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   228
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   229
lemma Says_to_Gets:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   230
     "Says A B X # evs \<in> set_cr ==> Gets B X # Says A B X # evs \<in> set_cr"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   231
by (rule set_cr.Reception, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   232
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   233
text{*The many nonces and keys generated, some simultaneously, force us to
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   234
  introduce them explicitly as shown below.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   235
lemma possibility_CR6:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   236
     "[|NC1 < (NC2::nat);  NC2 < NC3;  NC3 < NCA ;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   237
        NCA < NonceCCA;  NonceCCA < CardSecret;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   238
        KC1 < (KC2::key);  KC2 < KC3;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   239
        KC1 \<in> symKeys;  Key KC1 \<notin> used [];
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   240
        KC2 \<in> symKeys;  Key KC2 \<notin> used [];
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   241
        KC3 \<in> symKeys;  Key KC3 \<notin> used [];
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   242
        C = Cardholder k|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   243
   ==> \<exists>evs \<in> set_cr.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   244
       Says (CA i) C
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   245
            (Crypt KC2
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   246
             {|sign (priSK (CA i))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   247
                    {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   248
               certC (pan C) (pubSK (Cardholder k)) (XOR(CardSecret,NonceCCA))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   249
                     onlySig (priSK (CA i)),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   250
               cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   251
          \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   252
apply (intro exI bexI)
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   253
apply (rule_tac [2] 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   254
       set_cr.Nil 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   255
        [THEN set_cr.SET_CR1 [of concl: C i NC1], 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   256
         THEN Says_to_Gets, 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   257
	 THEN set_cr.SET_CR2 [of concl: i C NC1], 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   258
	 THEN Says_to_Gets,  
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   259
	 THEN set_cr.SET_CR3 [of concl: C i KC1 _ NC2], 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   260
	 THEN Says_to_Gets,  
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   261
	 THEN set_cr.SET_CR4 [of concl: i C NC2 NCA], 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   262
	 THEN Says_to_Gets,  
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   263
	 THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret],
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   264
	 THEN Says_to_Gets,  
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   265
	 THEN set_cr.SET_CR6 [of concl: i C KC2]])
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   266
apply (tactic "basic_possibility_tac")
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   267
apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   268
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   269
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   270
text{*General facts about message reception*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   271
lemma Gets_imp_Says:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   272
     "[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> \<exists>A. Says A B X \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   273
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   274
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   275
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   276
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   277
lemma Gets_imp_knows_Spy:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   278
     "[| Gets B X \<in> set evs; evs \<in> set_cr |]  ==> X \<in> knows Spy evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   279
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   280
declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   281
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   282
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   283
subsection{*Proofs on keys *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   284
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   285
text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   286
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   287
lemma Spy_see_private_Key [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   288
     "evs \<in> set_cr
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   289
      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   290
by (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   291
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   292
lemma Spy_analz_private_Key [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   293
     "evs \<in> set_cr ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   294
     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   295
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   296
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   297
declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   298
declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   299
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   300
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   301
subsection{*Begin Piero's Theorems on Certificates*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   302
text{*Trivial in the current model, where certificates by RCA are secure *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   303
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   304
lemma Crypt_valid_pubEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   305
     "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   306
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   307
         evs \<in> set_cr |] ==> EKi = pubEK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   308
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   309
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   310
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   311
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   312
lemma certificate_valid_pubEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   313
    "[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   314
        evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   315
     ==> EKi = pubEK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   316
apply (unfold cert_def signCert_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   317
apply (blast dest!: Crypt_valid_pubEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   318
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   319
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   320
lemma Crypt_valid_pubSK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   321
     "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   322
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   323
         evs \<in> set_cr |] ==> SKi = pubSK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   324
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   325
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   326
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   327
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   328
lemma certificate_valid_pubSK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   329
    "[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   330
        evs \<in> set_cr |] ==> SKi = pubSK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   331
apply (unfold cert_def signCert_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   332
apply (blast dest!: Crypt_valid_pubSK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   333
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   334
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   335
lemma Gets_certificate_valid:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   336
     "[| Gets A {| X, cert C EKi onlyEnc (priSK RCA),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   337
                      cert C SKi onlySig (priSK RCA)|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   338
         evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   339
      ==> EKi = pubEK C & SKi = pubSK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   340
by (blast dest: certificate_valid_pubEK certificate_valid_pubSK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   341
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   342
text{*Nobody can have used non-existent keys!*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   343
lemma new_keys_not_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   344
     "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   345
      ==> K \<notin> keysFor (parts (knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   346
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   347
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   348
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   349
apply (frule_tac [8] Gets_certificate_valid)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   350
apply (frule_tac [6] Gets_certificate_valid, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   351
txt{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   352
apply (force dest!: usedI keysFor_parts_insert)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   353
txt{*Others*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   354
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   355
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   356
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   357
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   358
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   359
subsection{*New versions: as above, but generalized to have the KK argument *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   360
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   361
lemma gen_new_keys_not_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   362
     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   363
      ==> Key K \<notin> used evs --> K \<in> symKeys -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   364
          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   365
by (auto simp add: new_keys_not_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   366
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   367
lemma gen_new_keys_not_analzd:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   368
     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   369
      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   370
by (blast intro: keysFor_mono [THEN [2] rev_subsetD]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   371
          dest: gen_new_keys_not_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   372
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   373
lemma analz_Key_image_insert_eq:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   374
     "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   375
      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   376
          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   377
by (simp add: gen_new_keys_not_analzd)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   378
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   379
lemma Crypt_parts_imp_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   380
     "[|Crypt K X \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   381
        K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   382
apply (rule ccontr)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   383
apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   384
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   385
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   386
lemma Crypt_analz_imp_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   387
     "[|Crypt K X \<in> analz (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   388
        K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   389
by (blast intro: Crypt_parts_imp_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   390
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   391
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   392
subsection{*Messages signed by CA*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   393
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   394
text{*Message @{text SET_CR2}: C can check CA's signature if he has received
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   395
     CA's certificate.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   396
lemma CA_Says_2_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   397
     "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC1|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   398
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   399
         evs \<in> set_cr; (CA i) \<notin> bad |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   400
     ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   401
                 \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   402
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   403
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   404
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   405
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   406
text{*Ever used?*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   407
lemma CA_Says_2:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   408
     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC1|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   409
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   410
         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   411
         evs \<in> set_cr; (CA i) \<notin> bad |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   412
      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   413
                  \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   414
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_2_lemma)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   415
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   416
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   417
text{*Message @{text SET_CR4}: C can check CA's signature if he has received
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   418
      CA's certificate.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   419
lemma CA_Says_4_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   420
     "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   421
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   422
         evs \<in> set_cr; (CA i) \<notin> bad |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   423
      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   424
                     {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   425
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   426
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   427
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   428
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   429
text{*NEVER USED*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   430
lemma CA_Says_4:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   431
     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   432
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   433
         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   434
         evs \<in> set_cr; (CA i) \<notin> bad |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   435
      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   436
                   {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   437
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_4_lemma)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   438
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   439
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   440
text{*Message @{text SET_CR6}: C can check CA's signature if he has
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   441
      received CA's certificate.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   442
lemma CA_Says_6_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   443
     "[| Crypt (priSK (CA i)) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   444
               (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   445
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   446
         evs \<in> set_cr; (CA i) \<notin> bad |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   447
      ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   448
      {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   449
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   450
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   451
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   452
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   453
text{*NEVER USED*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   454
lemma CA_Says_6:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   455
     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   456
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   457
         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   458
         evs \<in> set_cr; (CA i) \<notin> bad |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   459
      ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   460
                    {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   461
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_6_lemma)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   462
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   463
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   464
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   465
subsection{*Useful lemmas *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   466
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   467
text{*Rewriting rule for private encryption keys.  Analogous rewriting rules
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   468
for other keys aren't needed.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   469
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   470
lemma parts_image_priEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   471
     "[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   472
        evs \<in> set_cr|] ==> priEK C \<in> KK | C \<in> bad"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   473
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   474
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   475
text{*trivial proof because (priEK C) never appears even in (parts evs)*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   476
lemma analz_image_priEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   477
     "evs \<in> set_cr ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   478
          (Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   479
          (priEK C \<in> KK | C \<in> bad)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   480
by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   481
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   482
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   483
subsection{*Secrecy of Session Keys *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   484
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   485
subsubsection{*Lemmas about the predicate KeyCryptKey *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   486
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   487
text{*A fresh DK cannot be associated with any other
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   488
  (with respect to a given trace). *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   489
lemma DK_fresh_not_KeyCryptKey:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   490
     "[| Key DK \<notin> used evs; evs \<in> set_cr |] ==> ~ KeyCryptKey DK K evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   491
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   492
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   493
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   494
apply (blast dest: Crypt_analz_imp_used)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   495
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   496
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   497
text{*A fresh K cannot be associated with any other.  The assumption that
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   498
  DK isn't a private encryption key may be an artifact of the particular
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   499
  definition of KeyCryptKey.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   500
lemma K_fresh_not_KeyCryptKey:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   501
     "[|\<forall>C. DK \<noteq> priEK C; Key K \<notin> used evs|] ==> ~ KeyCryptKey DK K evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   502
apply (induct evs)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   503
apply (auto simp add: parts_insert2 split add: event.split)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   504
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   505
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   506
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   507
text{*This holds because if (priEK (CA i)) appears in any traffic then it must
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   508
  be known to the Spy, by @{term Spy_see_private_Key}*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   509
lemma cardSK_neq_priEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   510
     "[|Key cardSK \<notin> analz (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   511
        Key cardSK : parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   512
        evs \<in> set_cr|] ==> cardSK \<noteq> priEK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   513
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   514
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   515
lemma not_KeyCryptKey_cardSK [rule_format (no_asm)]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   516
     "[|cardSK \<notin> symKeys;  \<forall>C. cardSK \<noteq> priEK C;  evs \<in> set_cr|] ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   517
      Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptKey cardSK K evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   518
by (erule set_cr.induct, analz_mono_contra, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   519
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   520
text{*Lemma for message 5: pubSK C is never used to encrypt Keys.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   521
lemma pubSK_not_KeyCryptKey [simp]: "~ KeyCryptKey (pubSK C) K evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   522
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   523
apply (auto simp add: parts_insert2 split add: event.split)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   524
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   525
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   526
text{*Lemma for message 6: either cardSK is compromised (when we don't care)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   527
  or else cardSK hasn't been used to encrypt K.  Previously we treated
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   528
  message 5 in the same way, but the current model assumes that rule
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   529
  @{text SET_CR5} is executed only by honest agents.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   530
lemma msg6_KeyCryptKey_disj:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   531
     "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   532
          \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   533
        cardSK \<notin> symKeys;  evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   534
      ==> Key cardSK \<in> analz (knows Spy evs) |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   535
          (\<forall>K. ~ KeyCryptKey cardSK K evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   536
by (blast dest: not_KeyCryptKey_cardSK intro: cardSK_neq_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   537
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   538
text{*As usual: we express the property as a logical equivalence*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   539
lemma Key_analz_image_Key_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   540
     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K \<in> KK | Key K \<in> analz H)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   541
      ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   542
      P --> (Key K \<in> analz (Key`KK Un H)) = (K \<in> KK | Key K \<in> analz H)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   543
by (blast intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   544
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   545
ML
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   546
{*
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   547
val Gets_certificate_valid = thm "Gets_certificate_valid";
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   548
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   549
fun valid_certificate_tac i =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   550
    EVERY [ftac Gets_certificate_valid i,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   551
           assume_tac i,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   552
           etac conjE i, REPEAT (hyp_subst_tac i)];
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   553
*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   554
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   555
text{*The @{text "(no_asm)"} attribute is essential, since it retains
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   556
  the quantifier and allows the simprule's condition to itself be simplified.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   557
lemma symKey_compromise [rule_format (no_asm)]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   558
     "evs \<in> set_cr ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   559
      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. ~ KeyCryptKey K SK evs)   -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   560
               (Key SK \<in> analz (Key`KK Un (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   561
               (SK \<in> KK | Key SK \<in> analz (knows Spy evs)))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   562
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   563
apply (rule_tac [!] allI) +
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   564
apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   565
apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   566
apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   567
apply (erule_tac [9] msg6_KeyCryptKey_disj [THEN disjE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   568
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   569
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   570
         add: analz_image_keys_simps analz_knows_absorb
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   571
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   572
              K_fresh_not_KeyCryptKey
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   573
              DK_fresh_not_KeyCryptKey ball_conj_distrib
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   574
              analz_image_priEK disj_simps)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   575
  --{*46 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   576
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   577
txt{*3*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   578
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   579
txt{*5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   580
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   581
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   582
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   583
text{*The remaining quantifiers seem to be essential.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   584
  NO NEED to assume the cardholder's OK: bad cardholders don't do anything
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   585
  wrong!!*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   586
lemma symKey_secrecy [rule_format]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   587
     "[|CA i \<notin> bad;  K \<in> symKeys;  evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   588
      ==> \<forall>X c. Says (Cardholder c) (CA i) X \<in> set evs -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   589
                Key K \<in> parts{X} -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   590
                Cardholder c \<notin> bad -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   591
                Key K \<notin> analz (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   592
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   593
apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   594
apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   595
apply (erule_tac [11] msg6_KeyCryptKey_disj [THEN disjE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   596
apply (simp_all del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   597
         add: symKey_compromise fresh_notin_analz_knows_Spy
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   598
              analz_image_keys_simps analz_knows_absorb
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   599
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   600
              K_fresh_not_KeyCryptKey
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   601
              DK_fresh_not_KeyCryptKey
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   602
              analz_image_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   603
  --{*13 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   604
txt{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   605
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   606
apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   607
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   608
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   609
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   610
subsection{*Primary Goals of Cardholder Registration *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   611
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   612
text{*The cardholder's certificate really was created by the CA, provided the
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   613
    CA is uncompromised *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   614
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   615
text{*Lemma concerning the actual signed message digest*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   616
lemma cert_valid_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   617
     "[|Crypt (priSK (CA i)) {|Hash {|Nonce N, Pan(pan C)|}, Key cardSK, N1|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   618
          \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   619
        CA i \<notin> bad; evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   620
  ==> \<exists>KC2 X Y. Says (CA i) C
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   621
                     (Crypt KC2 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   622
                       {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   623
                  \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   624
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   625
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   626
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   627
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   628
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   629
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   630
text{*Pre-packaged version for cardholder.  We don't try to confirm the values
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   631
  of KC2, X and Y, since they are not important.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   632
lemma certificate_valid_cardSK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   633
    "[|Gets C (Crypt KC2 {|X, certC (pan C) cardSK N onlySig (invKey SKi),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   634
                              cert (CA i) SKi onlySig (priSK RCA)|}) \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   635
        CA i \<notin> bad; evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   636
  ==> \<exists>KC2 X Y. Says (CA i) C
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   637
                     (Crypt KC2 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   638
                       {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   639
                   \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   640
by (force dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Body]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   641
                    certificate_valid_pubSK cert_valid_lemma)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   642
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   643
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   644
lemma Hash_imp_parts [rule_format]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   645
     "evs \<in> set_cr
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   646
      ==> Hash{|X, Nonce N|} \<in> parts (knows Spy evs) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   647
          Nonce N \<in> parts (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   648
apply (erule set_cr.induct, force)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   649
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   650
apply (blast intro: parts_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   651
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   652
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   653
lemma Hash_imp_parts2 [rule_format]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   654
     "evs \<in> set_cr
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   655
      ==> Hash{|X, Nonce M, Y, Nonce N|} \<in> parts (knows Spy evs) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   656
          Nonce M \<in> parts (knows Spy evs) & Nonce N \<in> parts (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   657
apply (erule set_cr.induct, force)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   658
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   659
apply (blast intro: parts_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   660
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   661
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   662
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   663
subsection{*Secrecy of Nonces*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   664
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   665
subsubsection{*Lemmas about the predicate KeyCryptNonce *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   666
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   667
text{*A fresh DK cannot be associated with any other
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   668
  (with respect to a given trace). *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   669
lemma DK_fresh_not_KeyCryptNonce:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   670
     "[| DK \<in> symKeys; Key DK \<notin> used evs; evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   671
      ==> ~ KeyCryptNonce DK K evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   672
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   673
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   674
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   675
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   676
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   677
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   678
apply (auto simp add: DK_fresh_not_KeyCryptKey)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   679
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   680
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   681
text{*A fresh N cannot be associated with any other
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   682
      (with respect to a given trace). *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   683
lemma N_fresh_not_KeyCryptNonce:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   684
     "\<forall>C. DK \<noteq> priEK C ==> Nonce N \<notin> used evs --> ~ KeyCryptNonce DK N evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   685
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   686
apply (case_tac [2] "a")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   687
apply (auto simp add: parts_insert2)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   688
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   689
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   690
lemma not_KeyCryptNonce_cardSK [rule_format (no_asm)]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   691
     "[|cardSK \<notin> symKeys;  \<forall>C. cardSK \<noteq> priEK C;  evs \<in> set_cr|] ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   692
      Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptNonce cardSK N evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   693
apply (erule set_cr.induct, analz_mono_contra, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   694
txt{*6*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   695
apply (blast dest: not_KeyCryptKey_cardSK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   696
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   697
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   698
subsubsection{*Lemmas for message 5 and 6:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   699
  either cardSK is compromised (when we don't care)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   700
  or else cardSK hasn't been used to encrypt K. *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   701
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   702
text{*Lemma for message 5: pubSK C is never used to encrypt Nonces.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   703
lemma pubSK_not_KeyCryptNonce [simp]: "~ KeyCryptNonce (pubSK C) N evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   704
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   705
apply (auto simp add: parts_insert2 split add: event.split)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   706
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   707
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   708
text{*Lemma for message 6: either cardSK is compromised (when we don't care)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   709
  or else cardSK hasn't been used to encrypt K.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   710
lemma msg6_KeyCryptNonce_disj:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   711
     "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   712
          \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   713
        cardSK \<notin> symKeys;  evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   714
      ==> Key cardSK \<in> analz (knows Spy evs) |
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   715
          ((\<forall>K. ~ KeyCryptKey cardSK K evs) &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   716
           (\<forall>N. ~ KeyCryptNonce cardSK N evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   717
by (blast dest: not_KeyCryptKey_cardSK not_KeyCryptNonce_cardSK
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   718
          intro: cardSK_neq_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   719
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   720
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   721
text{*As usual: we express the property as a logical equivalence*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   722
lemma Nonce_analz_image_Key_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   723
     "P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   724
      ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   725
by (blast intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   726
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   727
text{*The @{text "(no_asm)"} attribute is essential, since it retains
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   728
  the quantifier and allows the simprule's condition to itself be simplified.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   729
lemma Nonce_compromise [rule_format (no_asm)]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   730
     "evs \<in> set_cr ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   731
      (\<forall>N KK. (\<forall>K \<in> KK. ~ KeyCryptNonce K N evs)   -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   732
               (Nonce N \<in> analz (Key`KK Un (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   733
               (Nonce N \<in> analz (knows Spy evs)))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   734
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   735
apply (rule_tac [!] allI)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   736
apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   737
apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   738
apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   739
apply (frule_tac [11] msg6_KeyCryptNonce_disj)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   740
apply (erule_tac [13] disjE)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   741
apply (simp_all del: image_insert image_Un
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   742
         add: symKey_compromise
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   743
              analz_image_keys_simps analz_knows_absorb
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   744
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   745
              N_fresh_not_KeyCryptNonce
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   746
              DK_fresh_not_KeyCryptNonce K_fresh_not_KeyCryptKey
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   747
              ball_conj_distrib analz_image_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   748
  --{*71 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   749
txt{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   750
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   751
txt{*3*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   752
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   753
txt{*5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   754
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   755
txt{*6*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   756
txt{*cardSK compromised?*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   757
apply (force del: allE ballE impCE simp add: symKey_compromise)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   758
txt{*Simplify again--necessary because the previous simplification introduces
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   759
  some logical connectives*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   760
by (force del: allE ballE impCE
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   761
          simp del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   762
          simp add: analz_image_keys_simps symKey_compromise)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   763
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   764
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   765
subsection{*Secrecy of CardSecret: the Cardholder's secret*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   766
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   767
lemma NC2_not_CardSecret:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   768
     "[|Crypt EKj {|Key K, Pan p, Hash {|Agent D, Nonce N|}|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   769
          \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   770
        Key K \<notin> analz (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   771
        Nonce N \<notin> analz (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   772
       evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   773
      ==> Crypt EKi {|Key K', Pan p', Nonce N|} \<notin> parts (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   774
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   775
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   776
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   777
apply (erule set_cr.induct, analz_mono_contra, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   778
apply (blast dest: Hash_imp_parts)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   779
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   780
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   781
lemma KC2_secure_lemma [rule_format]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   782
     "[|U = Crypt KC3 {|Agent C, Nonce N, Key KC2, X|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   783
        U \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   784
        evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   785
  ==> Nonce N \<notin> analz (knows Spy evs) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   786
      (\<exists>k i W. Says (Cardholder k) (CA i) {|U,W|} \<in> set evs & 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   787
               Cardholder k \<notin> bad & CA i \<notin> bad)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   788
apply (erule_tac P = "U \<in> ?H" in rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   789
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   790
apply (tactic{*valid_certificate_tac 8*})  --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   791
apply (simp_all del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   792
         add: analz_image_keys_simps analz_knows_absorb
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   793
              analz_knows_absorb2 notin_image_iff)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   794
  --{*19 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   795
apply (simp_all (no_asm_simp)) --{*leaves 4 subgoals*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   796
apply (blast intro!: analz_insertI)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   797
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   798
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   799
lemma KC2_secrecy:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   800
     "[|Gets B {|Crypt K {|Agent C, Nonce N, Key KC2, X|}, Y|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   801
        Nonce N \<notin> analz (knows Spy evs);  KC2 \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   802
        evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   803
       ==> Key KC2 \<notin> analz (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   804
by (force dest!: refl [THEN KC2_secure_lemma] symKey_secrecy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   805
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   806
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   807
text{*Inductive version*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   808
lemma CardSecret_secrecy_lemma [rule_format]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   809
     "[|CA i \<notin> bad;  evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   810
      ==> Key K \<notin> analz (knows Spy evs) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   811
          Crypt (pubEK (CA i)) {|Key K, Pan p, Nonce CardSecret|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   812
             \<in> parts (knows Spy evs) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   813
          Nonce CardSecret \<notin> analz (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   814
apply (erule set_cr.induct, analz_mono_contra)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   815
apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   816
apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   817
apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   818
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   819
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   820
         add: analz_image_keys_simps analz_knows_absorb
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   821
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   822
              EXHcrypt_def Crypt_notin_image_Key
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   823
              N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   824
              ball_conj_distrib Nonce_compromise symKey_compromise
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   825
              analz_image_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   826
  --{*12 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   827
txt{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   828
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   829
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   830
txt{*1*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   831
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   832
txt{*2*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   833
apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   834
txt{*3*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   835
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   836
txt{*4*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   837
apply (blast dest: NC2_not_CardSecret Gets_imp_knows_Spy [THEN analz.Inj] analz_symKeys_Decrypt)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   838
txt{*5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   839
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   840
txt{*6*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   841
apply (blast dest: KC2_secrecy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   842
apply (blast dest: KC2_secrecy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   843
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   844
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   845
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   846
text{*Packaged version for cardholder*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   847
lemma CardSecret_secrecy:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   848
     "[|Cardholder k \<notin> bad;  CA i \<notin> bad;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   849
        Says (Cardholder k) (CA i)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   850
           {|X, Crypt EKi {|Key KC3, Pan p, Nonce CardSecret|}|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   851
        Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   852
                    cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   853
        KC3 \<in> symKeys;  evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   854
      ==> Nonce CardSecret \<notin> analz (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   855
apply (frule Gets_certificate_valid, assumption)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   856
apply (subgoal_tac "Key KC3 \<notin> analz (knows Spy evs) ")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   857
apply (blast dest: CardSecret_secrecy_lemma)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   858
apply (rule symKey_secrecy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   859
apply (auto simp add: parts_insert2)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   860
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   861
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   862
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   863
subsection{*Secrecy of NonceCCA [the CA's secret] *}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   864
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   865
lemma NC2_not_NonceCCA:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   866
     "[|Hash {|Agent C', Nonce N', Agent C, Nonce N|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   867
          \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   868
        Nonce N \<notin> analz (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   869
       evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   870
      ==> Crypt KC1 {|{|Agent B, Nonce N|}, Hash p|} \<notin> parts (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   871
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   872
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   873
apply (erule set_cr.induct, analz_mono_contra, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   874
apply (blast dest: Hash_imp_parts2)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   875
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   876
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   877
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   878
text{*Inductive version*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   879
lemma NonceCCA_secrecy_lemma [rule_format]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   880
     "[|CA i \<notin> bad;  evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   881
      ==> Key K \<notin> analz (knows Spy evs) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   882
          Crypt K
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   883
            {|sign (priSK (CA i))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   884
                   {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   885
              X, Y|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   886
             \<in> parts (knows Spy evs) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   887
          Nonce NonceCCA \<notin> analz (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   888
apply (erule set_cr.induct, analz_mono_contra)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   889
apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   890
apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   891
apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   892
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   893
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   894
         add: analz_image_keys_simps analz_knows_absorb sign_def
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   895
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   896
              EXHcrypt_def Crypt_notin_image_Key
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   897
              N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   898
              ball_conj_distrib Nonce_compromise symKey_compromise
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   899
              analz_image_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   900
  --{*15 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   901
txt{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   902
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   903
txt{*1*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   904
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   905
txt{*2*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   906
apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   907
txt{*3*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   908
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   909
txt{*4*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   910
apply (blast dest: NC2_not_NonceCCA)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   911
txt{*5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   912
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   913
txt{*6*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   914
apply (blast dest: KC2_secrecy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   915
apply (blast dest: KC2_secrecy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   916
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   917
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   918
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   919
text{*Packaged version for cardholder*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   920
lemma NonceCCA_secrecy:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   921
     "[|Cardholder k \<notin> bad;  CA i \<notin> bad;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   922
        Gets (Cardholder k)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   923
           (Crypt KC2
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   924
            {|sign (priSK (CA i)) {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   925
              X, Y|}) \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   926
        Says (Cardholder k) (CA i)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   927
           {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, X'|}, Y'|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   928
        Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   929
                    cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   930
        KC2 \<in> symKeys;  evs \<in> set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   931
      ==> Nonce NonceCCA \<notin> analz (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   932
apply (frule Gets_certificate_valid, assumption)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   933
apply (subgoal_tac "Key KC2 \<notin> analz (knows Spy evs) ")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   934
apply (blast dest: NonceCCA_secrecy_lemma)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   935
apply (rule symKey_secrecy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   936
apply (auto simp add: parts_insert2)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   937
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   938
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   939
text{*We don't bother to prove guarantees for the CA.  He doesn't care about
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   940
  the PANSecret: it isn't his credit card!*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   941
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   942
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   943
subsection{*Rewriting Rule for PANs*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   944
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   945
text{*Lemma for message 6: either cardSK isn't a CA's private encryption key,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   946
  or if it is then (because it appears in traffic) that CA is bad,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   947
  and so the Spy knows that key already.  Either way, we can simplify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   948
  the expression @{term "analz (insert (Key cardSK) X)"}.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   949
lemma msg6_cardSK_disj:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   950
     "[|Gets A {|Crypt K {|c, n, k', Key cardSK, X|}, Y|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   951
          \<in> set evs;  evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   952
      ==> cardSK \<notin> range(invKey o pubEK o CA) | Key cardSK \<in> knows Spy evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   953
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   954
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   955
lemma analz_image_pan_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   956
     "(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H)  ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   957
      (Pan P \<in> analz (Key`nE Un H)) =   (Pan P \<in> analz H)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   958
by (blast intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   959
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   960
lemma analz_image_pan [rule_format]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   961
     "evs \<in> set_cr ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   962
       \<forall>KK. KK <= - invKey ` pubEK ` range CA -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   963
            (Pan P \<in> analz (Key`KK Un (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   964
            (Pan P \<in> analz (knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   965
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   966
apply (rule_tac [!] allI impI)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   967
apply (rule_tac [!] analz_image_pan_lemma)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   968
apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   969
apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   970
apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   971
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   972
         del: image_insert image_Un
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   973
         add: analz_image_keys_simps disjoint_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   974
              notin_image_iff analz_image_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   975
  --{*33 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   976
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   977
txt{*6*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   978
apply (simp add: insert_absorb)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   979
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   980
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   981
lemma analz_insert_pan:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   982
     "[| evs \<in> set_cr;  K \<notin> invKey ` pubEK ` range CA |] ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   983
          (Pan P \<in> analz (insert (Key K) (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   984
          (Pan P \<in> analz (knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   985
by (simp del: image_insert image_Un
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   986
         add: analz_image_keys_simps analz_image_pan)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   987
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   988
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   989
text{*Confidentiality of the PAN\@.  Maybe we could combine the statements of
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   990
  this theorem with @{term analz_image_pan}, requiring a single induction but
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   991
  a much more difficult proof.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   992
lemma pan_confidentiality:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   993
     "[| Pan (pan C) \<in> analz(knows Spy evs); C \<noteq>Spy; evs :set_cr|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   994
    ==> \<exists>i X K HN.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   995
        Says C (CA i) {|X, Crypt (pubEK (CA i)) {|Key K, Pan (pan C), HN|} |}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   996
           \<in> set evs
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   997
      & (CA i) \<in> bad"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   998
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   999
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1000
apply (tactic{*valid_certificate_tac 8*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1001
apply (tactic{*valid_certificate_tac 6*}) --{*for message 5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1002
apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1003
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1004
         del: image_insert image_Un
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1005
         add: analz_image_keys_simps analz_insert_pan analz_image_pan
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1006
              notin_image_iff analz_image_priEK)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1007
  --{*18 seconds on a 1.8GHz machine*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1008
txt{*fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1009
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1010
txt{*3*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1011
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1012
txt{*5*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1013
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1014
txt{*6*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1015
apply (simp (no_asm_simp) add: insert_absorb)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1016
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1017
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1018
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1019
subsection{*Unicity*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1020
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1021
lemma CR6_Says_imp_Notes:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1022
     "[|Says (CA i) C (Crypt KC2
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1023
          {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1024
            certC (pan C) cardSK X onlySig (priSK (CA i)),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1025
            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})  \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1026
        evs \<in> set_cr |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1027
      ==> Notes (CA i) (Key cardSK) \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1028
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1029
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1030
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1031
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1032
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1033
text{*Unicity of cardSK: it uniquely identifies the other components.  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1034
      This holds because a CA accepts a cardSK at most once.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1035
lemma cardholder_key_unicity:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1036
     "[|Says (CA i) C (Crypt KC2
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1037
          {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1038
            certC (pan C) cardSK X onlySig (priSK (CA i)),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1039
            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1040
          \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1041
        Says (CA i) C' (Crypt KC2'
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1042
          {|sign (priSK (CA i)) {|Agent C', Nonce NC3', Agent (CA i), Nonce Y'|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1043
            certC (pan C') cardSK X' onlySig (priSK (CA i)),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1044
            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1045
          \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1046
        evs \<in> set_cr |] ==> C=C' & NC3=NC3' & X=X' & KC2=KC2' & Y=Y'"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1047
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1048
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1049
apply (erule set_cr.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1050
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1051
apply (blast dest!: CR6_Says_imp_Notes)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1052
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1053
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1054
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1055
text{*UNUSED unicity result*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1056
lemma unique_KC1:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1057
     "[|Says C B {|Crypt KC1 X, Crypt EK {|Key KC1, Y|}|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1058
          \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1059
        Says C B' {|Crypt KC1 X', Crypt EK' {|Key KC1, Y'|}|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1060
          \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1061
        C \<notin> bad;  evs \<in> set_cr|] ==> B'=B & Y'=Y"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1062
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1063
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1064
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1065
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1066
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1067
text{*UNUSED unicity result*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1068
lemma unique_KC2:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1069
     "[|Says C B {|Crypt K {|Agent C, nn, Key KC2, X|}, Y|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1070
        Says C B' {|Crypt K' {|Agent C, nn', Key KC2, X'|}, Y'|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1071
        C \<notin> bad;  evs \<in> set_cr|] ==> B'=B & X'=X"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1072
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1073
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1074
apply (erule set_cr.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1075
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1076
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1077
text{*Cannot show cardSK to be secret because it isn't assumed to be fresh
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1078
  it could be a previously compromised cardSK [e.g. involving a bad CA]*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1079
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1080
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1081
end