src/HOL/SET_Protocol/Purchase.thy
author blanchet
Thu, 28 Aug 2014 16:58:27 +0200
changeset 58078 d44c9dc4bf30
parent 57492 74bf65a1910a
child 58889 5b7a9633cfa8
permissions -rw-r--r--
took out one more occurrence of 'PolyML.makestring'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33028
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
     1
(*  Title:      HOL/SET_Protocol/Purchase.thy
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
     2
    Author:     Giampaolo Bella
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
     3
    Author:     Fabio Massacci
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
     4
    Author:     Lawrence C Paulson
14199
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{*Purchase Phase of SET*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     8
33028
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
     9
theory Purchase
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
    10
imports Public_SET
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
    11
begin
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    12
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    13
text{*
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    14
Note: nonces seem to consist of 20 bytes.  That includes both freshness
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    15
challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    16
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    17
This version omits @{text LID_C} but retains @{text LID_M}. At first glance
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    18
(Programmer's Guide page 267) it seems that both numbers are just introduced
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    19
for the respective convenience of the Cardholder's and Merchant's
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    20
system. However, omitting both of them would create a problem of
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    21
identification: how can the Merchant's system know what transaction is it
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    22
supposed to process?
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    23
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    24
Further reading (Programmer's guide page 309) suggest that there is an outside
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    25
bootstrapping message (SET initiation message) which is used by the Merchant
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    26
and the Cardholder to agree on the actual transaction. This bootstrapping
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    27
message is described in the SET External Interface Guide and ought to generate
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    28
@{text LID_M}. According SET Extern Interface Guide, this number might be a
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    29
cookie, an invoice number etc. The Programmer's Guide on page 310, states that
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    30
in absence of @{text LID_M} the protocol must somehow ("outside SET") identify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    31
the transaction from OrderDesc, which is assumed to be a searchable text only
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    32
field. Thus, it is assumed that the Merchant or the Cardholder somehow agreed
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    33
out-of-bad on the value of @{text LID_M} (for instance a cookie in a web
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    34
transaction etc.). This out-of-band agreement is expressed with a preliminary
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    35
start action in which the merchant and the Cardholder agree on the appropriate
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    36
values. Agreed values are stored with a suitable notes action.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    37
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    38
"XID is a transaction ID that is usually generated by the Merchant system,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    39
unless there is no PInitRes, in which case it is generated by the Cardholder
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    40
system. It is a randomly generated 20 byte variable that is globally unique
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    41
(statistically). Merchant and Cardholder systems shall use appropriate random
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    42
number generators to ensure the global uniqueness of XID."
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    43
--Programmer's Guide, page 267.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    44
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    45
PI (Payment Instruction) is the most central and sensitive data structure in
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    46
SET. It is used to pass the data required to authorize a payment card payment
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    47
from the Cardholder to the Payment Gateway, which will use the data to
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    48
initiate a payment card transaction through the traditional payment card
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    49
financial network. The data is encrypted by the Cardholder and sent via the
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    50
Merchant, such that the data is hidden from the Merchant unless the Acquirer
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    51
passes the data back to the Merchant.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    52
--Programmer's Guide, page 271.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    53
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    54
consts
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    55
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    56
    CardSecret :: "nat => nat"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
    57
     --{*Maps Cardholders to CardSecrets.
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
    58
         A CardSecret of 0 means no cerificate, must use unsigned format.*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    59
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    60
    PANSecret :: "nat => nat"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
    61
     --{*Maps Cardholders to PANSecrets.*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    62
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
    63
inductive_set
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
    64
  set_pur :: "event list set"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
    65
where
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    66
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    67
  Nil:   --{*Initial trace is empty*}
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    68
         "[] \<in> set_pur"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    69
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
    70
| Fake:  --{*The spy MAY say anything he CAN say.*}
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    71
         "[| evsf \<in> set_pur;  X \<in> synth(analz(knows Spy evsf)) |]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    72
          ==> Says Spy B X  # evsf \<in> set_pur"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    73
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    74
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
    75
| Reception: --{*If A sends a message X to B, then B might receive it*}
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    76
             "[| evsr \<in> set_pur;  Says A B X \<in> set evsr |]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    77
              ==> Gets B X  # evsr \<in> set_pur"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    78
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
    79
| Start: 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    80
      --{*Added start event which is out-of-band for SET: the Cardholder and
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    81
          the merchant agree on the amounts and uses @{text LID_M} as an
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    82
          identifier.
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    83
          This is suggested by the External Interface Guide. The Programmer's
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    84
          Guide, in absence of @{text LID_M}, states that the merchant uniquely
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
    85
          identifies the order out of some data contained in OrderDesc.*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    86
   "[|evsStart \<in> set_pur;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    87
      Number LID_M \<notin> used evsStart;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    88
      C = Cardholder k; M = Merchant i; P = PG j;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    89
      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    90
      LID_M \<notin> range CardSecret;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    91
      LID_M \<notin> range PANSecret |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    92
     ==> Notes C {|Number LID_M, Transaction|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    93
       # Notes M {|Number LID_M, Agent P, Transaction|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    94
       # evsStart \<in> set_pur"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    95
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
    96
| PInitReq:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    97
     --{*Purchase initialization, page 72 of Formal Protocol Desc.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    98
   "[|evsPIReq \<in> set_pur;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    99
      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   100
      Nonce Chall_C \<notin> used evsPIReq;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   101
      Chall_C \<notin> range CardSecret; Chall_C \<notin> range PANSecret;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   102
      Notes C {|Number LID_M, Transaction |} \<in> set evsPIReq |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   103
    ==> Says C M {|Number LID_M, Nonce Chall_C|} # evsPIReq \<in> set_pur"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   104
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   105
| PInitRes:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   106
     --{*Merchant replies with his own label XID and the encryption
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   107
         key certificate of his chosen Payment Gateway. Page 74 of Formal
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   108
         Protocol Desc. We use @{text LID_M} to identify Cardholder*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   109
   "[|evsPIRes \<in> set_pur;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   110
      Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPIRes;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   111
      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   112
      Notes M {|Number LID_M, Agent P, Transaction|} \<in> set evsPIRes;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   113
      Nonce Chall_M \<notin> used evsPIRes;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   114
      Chall_M \<notin> range CardSecret; Chall_M \<notin> range PANSecret;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   115
      Number XID \<notin> used evsPIRes;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   116
      XID \<notin> range CardSecret; XID \<notin> range PANSecret|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   117
    ==> Says M C (sign (priSK M)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   118
                       {|Number LID_M, Number XID,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   119
                         Nonce Chall_C, Nonce Chall_M,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   120
                         cert P (pubEK P) onlyEnc (priSK RCA)|})
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   121
          # evsPIRes \<in> set_pur"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   122
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   123
| PReqUns:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   124
      --{*UNSIGNED Purchase request (CardSecret = 0).
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   125
        Page 79 of Formal Protocol Desc.
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   126
        Merchant never sees the amount in clear. This holds of the real
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   127
        protocol, where XID identifies the transaction. We omit
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   128
        Hash{|Number XID, Nonce (CardSecret k)|} from PIHead because
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   129
        the CardSecret is 0 and because AuthReq treated the unsigned case
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   130
        very differently from the signed one anyway.*}
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   131
   "!!Chall_C Chall_M OrderDesc P PurchAmt XID evsPReqU.
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   132
    [|evsPReqU \<in> set_pur;
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   133
      C = Cardholder k; CardSecret k = 0;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   134
      Key KC1 \<notin> used evsPReqU;  KC1 \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   135
      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   136
      HOD = Hash{|Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   137
      OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,Nonce Chall_M|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   138
      PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   139
      Gets C (sign (priSK M)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   140
                   {|Number LID_M, Number XID,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   141
                     Nonce Chall_C, Nonce Chall_M,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   142
                     cert P EKj onlyEnc (priSK RCA)|})
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   143
        \<in> set evsPReqU;
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   144
      Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqU;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   145
      Notes C {|Number LID_M, Transaction|} \<in> set evsPReqU |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   146
    ==> Says C M
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   147
             {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   148
               OIData, Hash{|PIHead, Pan (pan C)|} |}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   149
          # Notes C {|Key KC1, Agent M|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   150
          # evsPReqU \<in> set_pur"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   151
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   152
| PReqS:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   153
      --{*SIGNED Purchase request.  Page 77 of Formal Protocol Desc.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   154
          We could specify the equation
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   155
          @{term "PIReqSigned = {| PIDualSigned, OIDualSigned |}"}, since the
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   156
          Formal Desc. gives PIHead the same format in the unsigned case.
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   157
          However, there's little point, as P treats the signed and 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   158
          unsigned cases differently.*}
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   159
   "!!C Chall_C Chall_M EKj HOD KC2 LID_M M OIData
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   160
      OIDualSigned OrderDesc P PANData PIData PIDualSigned
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   161
      PIHead PurchAmt Transaction XID evsPReqS k.
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   162
    [|evsPReqS \<in> set_pur;
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   163
      C = Cardholder k;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   164
      CardSecret k \<noteq> 0;  Key KC2 \<notin> used evsPReqS;  KC2 \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   165
      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   166
      HOD = Hash{|Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   167
      OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD, Nonce Chall_M|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   168
      PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   169
                  Hash{|Number XID, Nonce (CardSecret k)|}|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   170
      PANData = {|Pan (pan C), Nonce (PANSecret k)|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   171
      PIData = {|PIHead, PANData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   172
      PIDualSigned = {|sign (priSK C) {|Hash PIData, Hash OIData|},
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   173
                       EXcrypt KC2 EKj {|PIHead, Hash OIData|} PANData|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   174
      OIDualSigned = {|OIData, Hash PIData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   175
      Gets C (sign (priSK M)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   176
                   {|Number LID_M, Number XID,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   177
                     Nonce Chall_C, Nonce Chall_M,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   178
                     cert P EKj onlyEnc (priSK RCA)|})
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   179
        \<in> set evsPReqS;
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   180
      Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqS;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   181
      Notes C {|Number LID_M, Transaction|} \<in> set evsPReqS |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   182
    ==> Says C M {|PIDualSigned, OIDualSigned|}
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   183
          # Notes C {|Key KC2, Agent M|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   184
          # evsPReqS \<in> set_pur"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   185
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   186
  --{*Authorization Request.  Page 92 of Formal Protocol Desc.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   187
    Sent in response to Purchase Request.*}
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   188
| AuthReq:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   189
   "[| evsAReq \<in> set_pur;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   190
       Key KM \<notin> used evsAReq;  KM \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   191
       Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   192
       HOD = Hash{|Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   193
       OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   194
                  Nonce Chall_M|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   195
       CardSecret k \<noteq> 0 -->
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   196
         P_I = {|sign (priSK C) {|HPIData, Hash OIData|}, encPANData|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   197
       Gets M {|P_I, OIData, HPIData|} \<in> set evsAReq;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   198
       Says M C (sign (priSK M) {|Number LID_M, Number XID,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   199
                                  Nonce Chall_C, Nonce Chall_M,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   200
                                  cert P EKj onlyEnc (priSK RCA)|})
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   201
         \<in> set evsAReq;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   202
        Notes M {|Number LID_M, Agent P, Transaction|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   203
           \<in> set evsAReq |]
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   204
    ==> Says M P
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   205
             (EncB (priSK M) KM (pubEK P)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   206
               {|Number LID_M, Number XID, Hash OIData, HOD|}   P_I)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   207
          # evsAReq \<in> set_pur"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   208
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   209
  --{*Authorization Response has two forms: for UNSIGNED and SIGNED PIs.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   210
    Page 99 of Formal Protocol Desc.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   211
    PI is a keyword (product!), so we call it @{text P_I}. The hashes HOD and
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   212
    HOIData occur independently in @{text P_I} and in M's message.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   213
    The authCode in AuthRes represents the baggage of EncB, which in the
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   214
    full protocol is [CapToken], [AcqCardMsg], [AuthToken]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   215
    optional items for split shipments, recurring payments, etc.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   216
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   217
| AuthResUns:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   218
    --{*Authorization Response, UNSIGNED*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   219
   "[| evsAResU \<in> set_pur;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   220
       C = Cardholder k; M = Merchant i;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   221
       Key KP \<notin> used evsAResU;  KP \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   222
       CardSecret k = 0;  KC1 \<in> symKeys;  KM \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   223
       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   224
       P_I = EXHcrypt KC1 EKj {|PIHead, HOIData|} (Pan (pan C));
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   225
       Gets P (EncB (priSK M) KM (pubEK P)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   226
               {|Number LID_M, Number XID, HOIData, HOD|} P_I)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   227
           \<in> set evsAResU |]
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   228
   ==> Says P M
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   229
            (EncB (priSK P) KP (pubEK M)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   230
              {|Number LID_M, Number XID, Number PurchAmt|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   231
              authCode)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   232
       # evsAResU \<in> set_pur"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   233
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   234
| AuthResS:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   235
    --{*Authorization Response, SIGNED*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   236
   "[| evsAResS \<in> set_pur;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   237
       C = Cardholder k;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   238
       Key KP \<notin> used evsAResS;  KP \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   239
       CardSecret k \<noteq> 0;  KC2 \<in> symKeys;  KM \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   240
       P_I = {|sign (priSK C) {|Hash PIData, HOIData|},
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   241
               EXcrypt KC2 (pubEK P) {|PIHead, HOIData|} PANData|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   242
       PANData = {|Pan (pan C), Nonce (PANSecret k)|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   243
       PIData = {|PIHead, PANData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   244
       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   245
                  Hash{|Number XID, Nonce (CardSecret k)|}|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   246
       Gets P (EncB (priSK M) KM (pubEK P)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   247
                {|Number LID_M, Number XID, HOIData, HOD|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   248
               P_I)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   249
           \<in> set evsAResS |]
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   250
   ==> Says P M
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   251
            (EncB (priSK P) KP (pubEK M)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   252
              {|Number LID_M, Number XID, Number PurchAmt|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   253
              authCode)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   254
       # evsAResS \<in> set_pur"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   255
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16474
diff changeset
   256
| PRes:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   257
    --{*Purchase response.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   258
   "[| evsPRes \<in> set_pur;  KP \<in> symKeys;  M = Merchant i;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   259
       Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   260
       Gets M (EncB (priSK P) KP (pubEK M)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   261
              {|Number LID_M, Number XID, Number PurchAmt|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   262
              authCode)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   263
          \<in> set evsPRes;
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   264
       Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPRes;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   265
       Says M P
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   266
            (EncB (priSK M) KM (pubEK P)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   267
              {|Number LID_M, Number XID, Hash OIData, HOD|} P_I)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   268
         \<in> set evsPRes;
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   269
       Notes M {|Number LID_M, Agent P, Transaction|}
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   270
          \<in> set evsPRes
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   271
      |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   272
   ==> Says M C
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   273
         (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   274
                           Hash (Number PurchAmt)|})
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   275
         # evsPRes \<in> set_pur"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   276
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   277
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   278
specification (CardSecret PANSecret)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   279
  inj_CardSecret:  "inj CardSecret"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   280
  inj_PANSecret:   "inj PANSecret"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   281
  CardSecret_neq_PANSecret: "CardSecret k \<noteq> PANSecret k'"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   282
    --{*No CardSecret equals any PANSecret*}
35703
29cb504abbb5 convert SET_Protocol to use Nat_Bijection library
huffman
parents: 33028
diff changeset
   283
  apply (rule_tac x="curry prod_encode 0" in exI)
29cb504abbb5 convert SET_Protocol to use Nat_Bijection library
huffman
parents: 33028
diff changeset
   284
  apply (rule_tac x="curry prod_encode 1" in exI)
29cb504abbb5 convert SET_Protocol to use Nat_Bijection library
huffman
parents: 33028
diff changeset
   285
  apply (simp add: prod_encode_eq inj_on_def)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   286
  done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   287
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   288
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   289
declare parts.Body [dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   290
declare analz_into_parts [dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   291
declare Fake_parts_insert_in_Un [dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   292
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   293
declare CardSecret_neq_PANSecret [iff] 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   294
        CardSecret_neq_PANSecret [THEN not_sym, iff]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   295
declare inj_CardSecret [THEN inj_eq, iff] 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   296
        inj_PANSecret [THEN inj_eq, iff]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   297
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   298
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   299
subsection{*Possibility Properties*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   300
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   301
lemma Says_to_Gets:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   302
     "Says A B X # evs \<in> set_pur ==> Gets B X # Says A B X # evs \<in> set_pur"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   303
by (rule set_pur.Reception, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   304
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   305
text{*Possibility for UNSIGNED purchases. Note that we need to ensure
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   306
that XID differs from OrderDesc and PurchAmt, since it is supposed to be
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   307
a unique number!*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   308
lemma possibility_Uns:
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   309
    "[| CardSecret k = 0;
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   310
        C = Cardholder k;  M = Merchant i;
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   311
        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   312
        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   313
        KC < KM; KM < KP;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   314
        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   315
        Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   316
        Chall_C < Chall_M; 
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   317
        Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   318
        Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   319
        LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   320
   ==> \<exists>evs \<in> set_pur.
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   321
          Says M C
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   322
               (sign (priSK M)
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   323
                    {|Number LID_M, Number XID, Nonce Chall_C, 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   324
                      Hash (Number PurchAmt)|})
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   325
                  \<in> set evs" 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   326
apply (intro exI bexI)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   327
apply (rule_tac [2]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   328
        set_pur.Nil
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   329
         [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   330
          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   331
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   332
          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   333
          THEN Says_to_Gets,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   334
          THEN set_pur.PReqUns [of concl: C M KC],
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   335
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   336
          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   337
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   338
          THEN set_pur.AuthResUns [of concl: "PG j" M KP LID_M XID],
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   339
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   340
          THEN set_pur.PRes]) 
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 30549
diff changeset
   341
apply basic_possibility
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   342
apply (simp_all add: used_Cons symKeys_neq_imp_neq) 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   343
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   344
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   345
lemma possibility_S:
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   346
    "[| CardSecret k \<noteq> 0;
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   347
        C = Cardholder k;  M = Merchant i;
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   348
        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   349
        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   350
        KC < KM; KM < KP;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   351
        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   352
        Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   353
        Chall_C < Chall_M; 
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   354
        Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   355
        Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   356
        LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   357
   ==>  \<exists>evs \<in> set_pur.
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   358
            Says M C
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   359
                 (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C, 
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   360
                                   Hash (Number PurchAmt)|})
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   361
               \<in> set evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   362
apply (intro exI bexI)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   363
apply (rule_tac [2]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   364
        set_pur.Nil
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   365
         [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   366
          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   367
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   368
          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   369
          THEN Says_to_Gets,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   370
          THEN set_pur.PReqS [of concl: C M _ _ KC],
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   371
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   372
          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   373
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   374
          THEN set_pur.AuthResS [of concl: "PG j" M KP LID_M XID],
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   375
          THEN Says_to_Gets, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   376
          THEN set_pur.PRes]) 
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 30549
diff changeset
   377
apply basic_possibility
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   378
apply (auto simp add: used_Cons symKeys_neq_imp_neq) 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   379
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   380
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   381
text{*General facts about message reception*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   382
lemma Gets_imp_Says:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   383
     "[| Gets B X \<in> set evs; evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   384
   ==> \<exists>A. Says A B X \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   385
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   386
apply (erule set_pur.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   387
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   388
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   389
lemma Gets_imp_knows_Spy:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   390
     "[| Gets B X \<in> set evs; evs \<in> set_pur |]  ==> X \<in> knows Spy evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   391
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   392
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   393
declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   394
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   395
text{*Forwarding lemmas, to aid simplification*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   396
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   397
lemma AuthReq_msg_in_parts_spies:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   398
     "[|Gets M {|P_I, OIData, HPIData|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   399
        evs \<in> set_pur|] ==> P_I \<in> parts (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   400
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   401
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   402
lemma AuthReq_msg_in_analz_spies:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   403
     "[|Gets M {|P_I, OIData, HPIData|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   404
        evs \<in> set_pur|] ==> P_I \<in> analz (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   405
by (blast dest: Gets_imp_knows_Spy [THEN analz.Inj])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   406
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   407
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   408
subsection{*Proofs on Asymmetric Keys*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   409
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   410
text{*Private Keys are Secret*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   411
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   412
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
   413
lemma Spy_see_private_Key [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   414
     "evs \<in> set_pur
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   415
      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   416
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   417
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   418
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   419
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   420
declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   421
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   422
lemma Spy_analz_private_Key [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   423
     "evs \<in> set_pur ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   424
     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   425
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   426
declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   427
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   428
text{*rewriting rule for priEK's*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   429
lemma parts_image_priEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   430
     "[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   431
        evs \<in> set_pur|] ==> priEK C \<in> KK | C \<in> bad"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   432
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   433
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   434
text{*trivial proof because @{term"priEK C"} never appears even in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   435
  @{term "parts evs"}. *}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   436
lemma analz_image_priEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   437
     "evs \<in> set_pur ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   438
          (Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   439
          (priEK C \<in> KK | C \<in> bad)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   440
by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   441
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   442
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   443
subsection{*Public Keys in Certificates are Correct*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   444
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   445
lemma Crypt_valid_pubEK [dest!]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   446
     "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   447
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   448
         evs \<in> set_pur |] ==> EKi = pubEK C"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   449
by (erule rev_mp, erule set_pur.induct, auto)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   450
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   451
lemma Crypt_valid_pubSK [dest!]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   452
     "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   453
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   454
         evs \<in> set_pur |] ==> SKi = pubSK C"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   455
by (erule rev_mp, erule set_pur.induct, auto)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   456
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   457
lemma certificate_valid_pubEK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   458
    "[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   459
        evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   460
     ==> EKi = pubEK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   461
by (unfold cert_def signCert_def, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   462
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   463
lemma certificate_valid_pubSK:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   464
    "[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   465
        evs \<in> set_pur |] ==> SKi = pubSK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   466
by (unfold cert_def signCert_def, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   467
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   468
lemma Says_certificate_valid [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   469
     "[| Says A B (sign SK {|lid, xid, cc, cm,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   470
                           cert C EK onlyEnc (priSK RCA)|}) \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   471
         evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   472
      ==> EK = pubEK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   473
by (unfold sign_def, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   474
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   475
lemma Gets_certificate_valid [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   476
     "[| Gets A (sign SK {|lid, xid, cc, cm,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   477
                           cert C EK onlyEnc (priSK RCA)|}) \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   478
         evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   479
      ==> EK = pubEK C"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   480
by (frule Gets_imp_Says, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   481
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   482
method_setup valid_certificate_tac = {*
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 24123
diff changeset
   483
  Args.goal_spec >> (fn quant =>
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 42814
diff changeset
   484
    fn ctxt => SIMPLE_METHOD'' quant (fn i =>
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 24123
diff changeset
   485
      EVERY [ftac @{thm Gets_certificate_valid} i,
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 42814
diff changeset
   486
             assume_tac i, REPEAT (hyp_subst_tac ctxt i)]))
42814
5af15f1e2ef6 simplified/unified method_setup/attribute_setup;
wenzelm
parents: 35703
diff changeset
   487
*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   488
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   489
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   490
subsection{*Proofs on Symmetric Keys*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   491
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   492
text{*Nobody can have used non-existent keys!*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   493
lemma new_keys_not_used [rule_format,simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   494
     "evs \<in> set_pur
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   495
      ==> Key K \<notin> used evs --> K \<in> symKeys -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   496
          K \<notin> keysFor (parts (knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   497
apply (erule set_pur.induct)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   498
apply (valid_certificate_tac [8]) --{*PReqS*}
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   499
apply (valid_certificate_tac [7]) --{*PReqUns*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   500
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   501
apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   502
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   503
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   504
lemma new_keys_not_analzd:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   505
     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   506
      ==> K \<notin> keysFor (analz (knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   507
by (blast intro: keysFor_mono [THEN [2] rev_subsetD] dest: new_keys_not_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   508
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   509
lemma Crypt_parts_imp_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   510
     "[|Crypt K X \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   511
        K \<in> symKeys; evs \<in> set_pur |] ==> Key K \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   512
apply (rule ccontr)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   513
apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   514
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   515
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   516
lemma Crypt_analz_imp_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   517
     "[|Crypt K X \<in> analz (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   518
        K \<in> symKeys; evs \<in> set_pur |] ==> Key K \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   519
by (blast intro: Crypt_parts_imp_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   520
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   521
text{*New versions: as above, but generalized to have the KK argument*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   522
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   523
lemma gen_new_keys_not_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   524
     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   525
      ==> Key K \<notin> used evs --> K \<in> symKeys -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   526
          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   527
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   528
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   529
lemma gen_new_keys_not_analzd:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   530
     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   531
      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   532
by (blast intro: keysFor_mono [THEN subsetD] dest: gen_new_keys_not_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   533
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   534
lemma analz_Key_image_insert_eq:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   535
     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   536
      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   537
          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   538
by (simp add: gen_new_keys_not_analzd)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   539
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   540
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   541
subsection{*Secrecy of Symmetric Keys*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   542
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   543
lemma Key_analz_image_Key_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   544
     "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
   545
      ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   546
      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
   547
by (blast intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   548
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   549
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   550
lemma symKey_compromise:
15205
ecf9a884970d some x-symbols
paulson
parents: 14256
diff changeset
   551
     "evs \<in> set_pur \<Longrightarrow>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   552
      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow>
15205
ecf9a884970d some x-symbols
paulson
parents: 14256
diff changeset
   553
        (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) \<longrightarrow>
ecf9a884970d some x-symbols
paulson
parents: 14256
diff changeset
   554
               (Key SK \<in> analz (Key`KK \<union> (knows Spy evs))) =
ecf9a884970d some x-symbols
paulson
parents: 14256
diff changeset
   555
               (SK \<in> KK \<or> Key SK \<in> analz (knows Spy evs)))"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   556
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   557
apply (rule_tac [!] allI)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   558
apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   559
apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   560
apply (valid_certificate_tac [8]) --{*PReqS*}
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   561
apply (valid_certificate_tac [7]) --{*PReqUns*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   562
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   563
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   564
         add: analz_image_keys_simps disj_simps
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   565
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   566
              analz_insert_simps analz_image_priEK)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   567
  --{*8 seconds on a 1.6GHz machine*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   568
apply spy_analz --{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   569
apply (blast elim!: ballE)+ --{*PReq: unsigned and signed*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   570
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   571
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   572
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   573
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   574
subsection{*Secrecy of Nonces*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   575
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   576
text{*As usual: we express the property as a logical equivalence*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   577
lemma Nonce_analz_image_Key_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   578
     "P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   579
      ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   580
by (blast intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   581
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   582
text{*The @{text "(no_asm)"} attribute is essential, since it retains
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   583
  the quantifier and allows the simprule's condition to itself be simplified.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   584
lemma Nonce_compromise [rule_format (no_asm)]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   585
     "evs \<in> set_pur ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   586
      (\<forall>N KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C))   -->
14256
paulson
parents: 14206
diff changeset
   587
              (Nonce N \<in> analz (Key`KK \<union> (knows Spy evs))) =
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   588
              (Nonce N \<in> analz (knows Spy evs)))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   589
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   590
apply (rule_tac [!] allI)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   591
apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   592
apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   593
apply (valid_certificate_tac [8]) --{*PReqS*}
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   594
apply (valid_certificate_tac [7]) --{*PReqUns*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   595
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   596
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   597
         add: analz_image_keys_simps disj_simps symKey_compromise
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   598
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   599
              analz_insert_simps analz_image_priEK)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   600
  --{*8 seconds on a 1.6GHz machine*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   601
apply spy_analz --{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   602
apply (blast elim!: ballE) --{*PReqS*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   603
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   604
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   605
lemma PANSecret_notin_spies:
14256
paulson
parents: 14206
diff changeset
   606
     "[|Nonce (PANSecret k) \<in> analz (knows Spy evs);  evs \<in> set_pur|]
paulson
parents: 14206
diff changeset
   607
      ==> 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   608
       (\<exists>V W X Y KC2 M. \<exists>P \<in> bad.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   609
          Says (Cardholder k) M
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   610
               {|{|W, EXcrypt KC2 (pubEK P) X {|Y, Nonce (PANSecret k)|}|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   611
                 V|}  \<in>  set evs)"
14256
paulson
parents: 14206
diff changeset
   612
apply (erule rev_mp)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   613
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   614
apply (frule_tac [9] AuthReq_msg_in_analz_spies)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   615
apply (valid_certificate_tac [8]) --{*PReqS*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   616
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   617
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   618
         add: analz_image_keys_simps disj_simps
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   619
              symKey_compromise pushes sign_def Nonce_compromise
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   620
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   621
              analz_insert_simps analz_image_priEK)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   622
  --{*2.5 seconds on a 1.6GHz machine*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   623
apply spy_analz
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   624
apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   625
apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj] 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   626
                   Gets_imp_knows_Spy [THEN analz.Inj])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   627
apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj]) --{*PReqS*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   628
apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj] 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   629
                   Gets_imp_knows_Spy [THEN analz.Inj]) --{*PRes*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   630
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   631
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   632
text{*This theorem is a bit silly, in that many CardSecrets are 0!
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   633
  But then we don't care.  NOT USED*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   634
lemma CardSecret_notin_spies:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   635
     "evs \<in> set_pur ==> Nonce (CardSecret i) \<notin> parts (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   636
by (erule set_pur.induct, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   637
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   638
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   639
subsection{*Confidentiality of PAN*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   640
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   641
lemma analz_image_pan_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   642
     "(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H)  ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   643
      (Pan P \<in> analz (Key`nE Un H)) =   (Pan P \<in> analz H)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   644
by (blast intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   645
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   646
text{*The @{text "(no_asm)"} attribute is essential, since it retains
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   647
  the quantifier and allows the simprule's condition to itself be simplified.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   648
lemma analz_image_pan [rule_format (no_asm)]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   649
     "evs \<in> set_pur ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   650
       \<forall>KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) -->
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   651
            (Pan P \<in> analz (Key`KK Un (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   652
            (Pan P \<in> analz (knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   653
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   654
apply (rule_tac [!] allI impI)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   655
apply (rule_tac [!] analz_image_pan_lemma)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   656
apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   657
apply (valid_certificate_tac [8]) --{*PReqS*}
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   658
apply (valid_certificate_tac [7]) --{*PReqUns*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   659
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   660
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   661
         add: analz_image_keys_simps
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   662
              symKey_compromise pushes sign_def
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   663
              analz_Key_image_insert_eq notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   664
              analz_insert_simps analz_image_priEK)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   665
  --{*7 seconds on a 1.6GHz machine*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   666
apply spy_analz --{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   667
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   668
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   669
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   670
lemma analz_insert_pan:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   671
     "[| evs \<in> set_pur;  K \<notin> range(\<lambda>C. priEK C) |] ==>
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   672
          (Pan P \<in> analz (insert (Key K) (knows Spy evs))) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   673
          (Pan P \<in> analz (knows Spy evs))"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   674
by (simp del: image_insert image_Un
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   675
         add: analz_image_keys_simps analz_image_pan)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   676
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   677
text{*Confidentiality of the PAN, unsigned case.*}
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   678
theorem pan_confidentiality_unsigned:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   679
     "[| Pan (pan C) \<in> analz(knows Spy evs);  C = Cardholder k;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   680
         CardSecret k = 0;  evs \<in> set_pur|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   681
    ==> \<exists>P M KC1 K X Y.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   682
     Says C M {|EXHcrypt KC1 (pubEK P) X (Pan (pan C)), Y|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   683
          \<in> set evs  &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   684
     P \<in> bad"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   685
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   686
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   687
apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   688
apply (valid_certificate_tac [8]) --{*PReqS*}
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   689
apply (valid_certificate_tac [7]) --{*PReqUns*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   690
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   691
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   692
         add: analz_image_keys_simps analz_insert_pan analz_image_pan
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   693
              notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   694
              analz_insert_simps analz_image_priEK)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   695
  --{*3 seconds on a 1.6GHz machine*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   696
apply spy_analz --{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   697
apply blast --{*PReqUns: unsigned*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   698
apply force --{*PReqS: signed*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   699
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   700
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   701
text{*Confidentiality of the PAN, signed case.*}
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   702
theorem pan_confidentiality_signed:
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   703
 "[|Pan (pan C) \<in> analz(knows Spy evs);  C = Cardholder k;
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   704
    CardSecret k \<noteq> 0;  evs \<in> set_pur|]
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   705
  ==> \<exists>P M KC2 PIDualSign_1 PIDualSign_2 other OIDualSign.
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   706
      Says C M {|{|PIDualSign_1, 
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   707
                   EXcrypt KC2 (pubEK P) PIDualSign_2 {|Pan (pan C), other|}|}, 
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   708
       OIDualSign|} \<in> set evs  &  P \<in> bad"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   709
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   710
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   711
apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   712
apply (valid_certificate_tac [8]) --{*PReqS*}
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   713
apply (valid_certificate_tac [7]) --{*PReqUns*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   714
apply (simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   715
         del: image_insert image_Un imp_disjL
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   716
         add: analz_image_keys_simps analz_insert_pan analz_image_pan
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   717
              notin_image_iff
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   718
              analz_insert_simps analz_image_priEK)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   719
  --{*3 seconds on a 1.6GHz machine*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   720
apply spy_analz --{*Fake*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   721
apply force --{*PReqUns: unsigned*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   722
apply blast --{*PReqS: signed*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   723
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   724
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   725
text{*General goal: that C, M and PG agree on those details of the transaction
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   726
     that they are allowed to know about.  PG knows about price and account
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   727
     details.  M knows about the order description and price.  C knows both.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   728
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   729
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   730
subsection{*Proofs Common to Signed and Unsigned Versions*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   731
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   732
lemma M_Notes_PG:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   733
     "[|Notes M {|Number LID_M, Agent P, Agent M, Agent C, etc|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   734
        evs \<in> set_pur|] ==> \<exists>j. P = PG j"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   735
by (erule rev_mp, erule set_pur.induct, simp_all)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   736
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   737
text{*If we trust M, then @{term LID_M} determines his choice of P
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   738
      (Payment Gateway)*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   739
lemma goodM_gives_correct_PG:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   740
     "[| MsgPInitRes = 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   741
            {|Number LID_M, xid, cc, cm, cert P EKj onlyEnc (priSK RCA)|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   742
         Crypt (priSK M) (Hash MsgPInitRes) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   743
         evs \<in> set_pur; M \<notin> bad |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   744
      ==> \<exists>j trans.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   745
            P = PG j &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   746
            Notes M {|Number LID_M, Agent P, trans|} \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   747
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   748
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   749
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   750
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   751
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   752
apply (blast intro: M_Notes_PG)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   753
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   754
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   755
lemma C_gets_correct_PG:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   756
     "[| Gets A (sign (priSK M) {|Number LID_M, xid, cc, cm,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   757
                              cert P EKj onlyEnc (priSK RCA)|}) \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   758
         evs \<in> set_pur;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   759
      ==> \<exists>j trans.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   760
            P = PG j &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   761
            Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   762
            EKj = pubEK P"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   763
by (rule refl [THEN goodM_gives_correct_PG, THEN exE], auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   764
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   765
text{*When C receives PInitRes, he learns M's choice of P*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   766
lemma C_verifies_PInitRes:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   767
 "[| MsgPInitRes = {|Number LID_M, Number XID, Nonce Chall_C, Nonce Chall_M,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   768
           cert P EKj onlyEnc (priSK RCA)|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   769
     Crypt (priSK M) (Hash MsgPInitRes) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   770
     evs \<in> set_pur;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   771
  ==> \<exists>j trans.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   772
         Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   773
         P = PG j &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   774
         EKj = pubEK P"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   775
apply clarify
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_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   778
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   779
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   780
apply (blast intro: M_Notes_PG)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   781
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   782
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   783
text{*Corollary of previous one*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   784
lemma Says_C_PInitRes:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   785
     "[|Says A C (sign (priSK M)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   786
                      {|Number LID_M, Number XID,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   787
                        Nonce Chall_C, Nonce Chall_M,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   788
                        cert P EKj onlyEnc (priSK RCA)|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   789
           \<in> set evs;  M \<notin> bad;  evs \<in> set_pur|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   790
      ==> \<exists>j trans.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   791
           Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   792
           P = PG j &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   793
           EKj = pubEK (PG j)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   794
apply (frule Says_certificate_valid)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   795
apply (auto simp add: sign_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   796
apply (blast dest: refl [THEN goodM_gives_correct_PG])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   797
apply (blast dest: refl [THEN C_verifies_PInitRes])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   798
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   799
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   800
text{*When P receives an AuthReq, he knows that the signed part originated 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   801
      with M. PIRes also has a signed message from M....*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   802
lemma P_verifies_AuthReq:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   803
     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   804
         Crypt (priSK M) (Hash {|AuthReqData, Hash P_I|})
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   805
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   806
         evs \<in> set_pur;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   807
      ==> \<exists>j trans KM OIData HPIData.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   808
            Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   809
            Gets M {|P_I, OIData, HPIData|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   810
            Says M (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   811
              \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   812
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   813
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   814
apply (erule set_pur.induct, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   815
apply (frule_tac [4] M_Notes_PG, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   816
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   817
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   818
text{*When M receives AuthRes, he knows that P signed it, including
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
   819
  the identifying tags and the purchase amount, which he can verify.
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   820
  (Although the spec has SIGNED and UNSIGNED forms of AuthRes, they
16474
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   821
   send the same message to M.)  The conclusion is weak: M is existentially
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   822
  quantified! That is because Authorization Response does not refer to M, while
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   823
  the digital envelope weakens the link between @{term MsgAuthRes} and
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   824
  @{term"priSK M"}.  Changing the precondition to refer to 
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   825
  @{term "Crypt K (sign SK M)"} requires assuming @{term K} to be secure, since
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   826
  otherwise the Spy could create that message.*}
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   827
theorem M_verifies_AuthRes:
16474
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   828
  "[| MsgAuthRes = {|{|Number LID_M, Number XID, Number PurchAmt|}, 
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   829
                     Hash authCode|};
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   830
      Crypt (priSK (PG j)) (Hash MsgAuthRes) \<in> parts (knows Spy evs);
16474
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   831
      PG j \<notin> bad;  evs \<in> set_pur|]
c3c0208988c0 fixed a faulty proof
paulson
parents: 16417
diff changeset
   832
   ==> \<exists>M KM KP HOIData HOD P_I.
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   833
        Gets (PG j)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   834
           (EncB (priSK M) KM (pubEK (PG j))
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   835
                    {|Number LID_M, Number XID, HOIData, HOD|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   836
                    P_I) \<in> set evs &
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   837
        Says (PG j) M
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   838
             (EncB (priSK (PG j)) KP (pubEK M)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   839
              {|Number LID_M, Number XID, Number PurchAmt|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
   840
              authCode) \<in> set evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   841
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   842
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   843
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   844
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   845
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   846
apply blast+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   847
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   848
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   849
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   850
subsection{*Proofs for Unsigned Purchases*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   851
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   852
text{*What we can derive from the ASSUMPTION that C issued a purchase request.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   853
   In the unsigned case, we must trust "C": there's no authentication.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   854
lemma C_determines_EKj:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   855
     "[| Says C M {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   856
                    OIData, Hash{|PIHead, Pan (pan C)|} |} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   857
         PIHead = {|Number LID_M, Trans_details|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   858
         evs \<in> set_pur;  C = Cardholder k;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   859
  ==> \<exists>trans j.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   860
               Notes M {|Number LID_M, Agent (PG j), trans |} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   861
               EKj = pubEK (PG j)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   862
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   863
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   864
apply (erule set_pur.induct, simp_all)
24123
a0fc58900606 tuned ML bindings (for multithreading);
wenzelm
parents: 23755
diff changeset
   865
apply (valid_certificate_tac [2]) --{*PReqUns*}
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   866
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   867
apply (blast dest: Gets_imp_Says Says_C_PInitRes)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   868
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   869
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   870
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   871
text{*Unicity of @{term LID_M} between Merchant and Cardholder notes*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   872
lemma unique_LID_M:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   873
     "[|Notes (Merchant i) {|Number LID_M, Agent P, Trans|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   874
        Notes C {|Number LID_M, Agent M, Agent C, Number OD,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   875
             Number PA|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   876
        evs \<in> set_pur|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   877
      ==> M = Merchant i & Trans = {|Agent M, Agent C, Number OD, Number PA|}"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   878
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   879
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   880
apply (erule set_pur.induct, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   881
apply (force dest!: Notes_imp_parts_subset_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   882
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   883
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   884
text{*Unicity of @{term LID_M}, for two Merchant Notes events*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   885
lemma unique_LID_M2:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   886
     "[|Notes M {|Number LID_M, Trans|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   887
        Notes M {|Number LID_M, Trans'|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   888
        evs \<in> set_pur|] ==> Trans' = Trans"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   889
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   890
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   891
apply (erule set_pur.induct, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   892
apply (force dest!: Notes_imp_parts_subset_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   893
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   894
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   895
text{*Lemma needed below: for the case that
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   896
  if PRes is present, then @{term LID_M} has been used.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   897
lemma signed_imp_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   898
     "[| Crypt (priSK M) (Hash X) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   899
         M \<notin> bad;  evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   900
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   901
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   902
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   903
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   904
apply safe
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   905
apply blast+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   906
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   907
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   908
text{*Similar, with nested Hash*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   909
lemma signed_Hash_imp_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   910
     "[| Crypt (priSK C) (Hash {|H, Hash X|}) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   911
         C \<notin> bad;  evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   912
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   913
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   914
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   915
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   916
apply safe
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   917
apply blast+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   918
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   919
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   920
text{*Lemma needed below: for the case that
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   921
  if PRes is present, then @{text LID_M} has been used.*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   922
lemma PRes_imp_LID_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   923
     "[| Crypt (priSK M) (Hash {|N, X|}) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   924
         M \<notin> bad;  evs \<in> set_pur|] ==> N \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   925
by (drule signed_imp_used, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   926
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   927
text{*When C receives PRes, he knows that M and P agreed to the purchase details.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   928
  He also knows that P is the same PG as before*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   929
lemma C_verifies_PRes_lemma:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   930
     "[| Crypt (priSK M) (Hash MsgPRes) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   931
         Notes C {|Number LID_M, Trans |} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   932
         Trans = {| Agent M, Agent C, Number OrderDesc, Number PurchAmt |};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   933
         MsgPRes = {|Number LID_M, Number XID, Nonce Chall_C,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   934
                Hash (Number PurchAmt)|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   935
         evs \<in> set_pur;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   936
  ==> \<exists>j KP.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   937
        Notes M {|Number LID_M, Agent (PG j), Trans |}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   938
          \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   939
        Gets M (EncB (priSK (PG j)) KP (pubEK M)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   940
                {|Number LID_M, Number XID, Number PurchAmt|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   941
                authCode)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   942
          \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   943
        Says M C (sign (priSK M) MsgPRes) \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   944
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   945
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   946
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   947
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   948
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   949
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   950
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   951
apply blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   952
apply (blast dest: PRes_imp_LID_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   953
apply (frule M_Notes_PG, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   954
apply (blast dest: unique_LID_M)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   955
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   956
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   957
text{*When the Cardholder receives Purchase Response from an uncompromised
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   958
Merchant, he knows that M sent it. He also knows that M received a message signed
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   959
by a Payment Gateway chosen by M to authorize the purchase.*}
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   960
theorem C_verifies_PRes:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   961
     "[| MsgPRes = {|Number LID_M, Number XID, Nonce Chall_C,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   962
                     Hash (Number PurchAmt)|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   963
         Gets C (sign (priSK M) MsgPRes) \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   964
         Notes C {|Number LID_M, Agent M, Agent C, Number OrderDesc,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   965
                   Number PurchAmt|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   966
         evs \<in> set_pur;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   967
  ==> \<exists>P KP trans.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   968
        Notes M {|Number LID_M,Agent P, trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   969
        Gets M (EncB (priSK P) KP (pubEK M)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   970
                {|Number LID_M, Number XID, Number PurchAmt|}
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
   971
                authCode)  \<in>  set evs &
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   972
        Says M C (sign (priSK M) MsgPRes) \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   973
apply (rule C_verifies_PRes_lemma [THEN exE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   974
apply (auto simp add: sign_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   975
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   976
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   977
subsection{*Proofs for Signed Purchases*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   978
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   979
text{*Some Useful Lemmas: the cardholder knows what he is doing*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   980
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   981
lemma Crypt_imp_Says_Cardholder:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   982
     "[| Crypt K {|{|{|Number LID_M, others|}, Hash OIData|}, Hash PANData|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   983
           \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   984
         PANData = {|Pan (pan (Cardholder k)), Nonce (PANSecret k)|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   985
         Key K \<notin> analz (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   986
         evs \<in> set_pur|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   987
  ==> \<exists>M shash EK HPIData.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   988
       Says (Cardholder k) M {|{|shash,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   989
          Crypt K
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   990
            {|{|{|Number LID_M, others|}, Hash OIData|}, Hash PANData|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   991
           Crypt EK {|Key K, PANData|}|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   992
          OIData, HPIData|} \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   993
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   994
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   995
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   996
apply (erule set_pur.induct, analz_mono_contra)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   997
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   998
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   999
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1000
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1001
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1002
lemma Says_PReqS_imp_trans_details_C:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1003
     "[| MsgPReqS = {|{|shash,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1004
                 Crypt K
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1005
                  {|{|{|Number LID_M, PIrest|}, Hash OIData|}, hashpd|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1006
            cryptek|}, data|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1007
         Says (Cardholder k) M MsgPReqS \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1008
         evs \<in> set_pur |]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1009
   ==> \<exists>trans.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1010
           Notes (Cardholder k) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1011
                 {|Number LID_M, Agent M, Agent (Cardholder k), trans|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1012
            \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1013
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1014
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1015
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1016
apply (simp_all (no_asm_simp))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1017
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1018
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1019
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1020
text{*Can't happen: only Merchants create this type of Note*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1021
lemma Notes_Cardholder_self_False:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1022
     "[|Notes (Cardholder k)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1023
          {|Number n, Agent P, Agent (Cardholder k), Agent C, etc|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1024
        evs \<in> set_pur|] ==> False"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15214
diff changeset
  1025
by (erule rev_mp, erule set_pur.induct, auto)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1026
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1027
text{*When M sees a dual signature, he knows that it originated with C.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1028
  Using XID he knows it was intended for him.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1029
  This guarantee isn't useful to P, who never gets OIData.*}
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1030
theorem M_verifies_Signed_PReq:
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1031
 "[| MsgDualSign = {|HPIData, Hash OIData|};
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1032
     OIData = {|Number LID_M, etc|};
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1033
     Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1034
     Notes M {|Number LID_M, Agent P, extras|} \<in> set evs;
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1035
     M = Merchant i;  C = Cardholder k;  C \<notin> bad;  evs \<in> set_pur|]
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1036
  ==> \<exists>PIData PICrypt.
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
  1037
        HPIData = Hash PIData &
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
  1038
        Says C M {|{|sign (priSK C) MsgDualSign, PICrypt|}, OIData, Hash PIData|}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32404
diff changeset
  1039
          \<in> set evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1040
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1041
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1042
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1043
apply (erule set_pur.induct)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1044
apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1045
apply simp_all
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1046
apply blast
32404
da3ca3c6ec81 sledgehammer used to streamline protocol proofs
paulson
parents: 30607
diff changeset
  1047
apply (metis subsetD insert_subset parts.Fst parts_increasing signed_Hash_imp_used)
da3ca3c6ec81 sledgehammer used to streamline protocol proofs
paulson
parents: 30607
diff changeset
  1048
apply (metis unique_LID_M)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1049
apply (blast dest!: Notes_Cardholder_self_False)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1050
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1051
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1052
text{*When P sees a dual signature, he knows that it originated with C.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1053
  and was intended for M. This guarantee isn't useful to M, who never gets
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1054
  PIData. I don't see how to link @{term "PG j"} and @{text LID_M} without
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1055
  assuming @{term "M \<notin> bad"}.*}
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1056
theorem P_verifies_Signed_PReq:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1057
     "[| MsgDualSign = {|Hash PIData, HOIData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1058
         PIData = {|PIHead, PANData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1059
         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1060
                    TransStain|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1061
         Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1062
         evs \<in> set_pur;  C \<notin> bad;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1063
    ==> \<exists>OIData OrderDesc K j trans.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1064
          HOD = Hash{|Number OrderDesc, Number PurchAmt|} &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1065
          HOIData = Hash OIData &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1066
          Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1067
          Says C M {|{|sign (priSK C) MsgDualSign,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1068
                     EXcrypt K (pubEK (PG j))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1069
                                {|PIHead, Hash OIData|} PANData|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1070
                     OIData, Hash PIData|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1071
            \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1072
apply clarify
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_pur.induct, simp_all)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1075
apply (auto dest!: C_gets_correct_PG)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1076
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1077
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1078
lemma C_determines_EKj_signed:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1079
     "[| Says C M {|{|sign (priSK C) text,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1080
                      EXcrypt K EKj {|PIHead, X|} Y|}, Z|} \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1081
         PIHead = {|Number LID_M, Number XID, W|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1082
         C = Cardholder k;  evs \<in> set_pur;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1083
  ==> \<exists> trans j.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1084
         Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1085
         EKj = pubEK (PG j)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1086
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1087
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1088
apply (erule set_pur.induct, simp_all, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1089
apply (blast dest: C_gets_correct_PG)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1090
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1091
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1092
lemma M_Says_AuthReq:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1093
     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1094
         sign (priSK M) {|AuthReqData, Hash P_I|} \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1095
         evs \<in> set_pur;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1096
   ==> \<exists>j trans KM.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1097
           Notes M {|Number LID_M, Agent (PG j), trans |} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1098
             Says M (PG j)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1099
               (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1100
              \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1101
apply (rule refl [THEN P_verifies_AuthReq, THEN exE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1102
apply (auto simp add: sign_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1103
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1104
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1105
text{*A variant of @{text M_verifies_Signed_PReq} with explicit PI information.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1106
  Even here we cannot be certain about what C sent to M, since a bad
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1107
  PG could have replaced the two key fields.  (NOT USED)*}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1108
lemma Signed_PReq_imp_Says_Cardholder:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1109
     "[| MsgDualSign = {|Hash PIData, Hash OIData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1110
         OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD, etc|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1111
         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1112
                    TransStain|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1113
         PIData = {|PIHead, PANData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1114
         Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1115
         M = Merchant i;  C = Cardholder k;  C \<notin> bad;  evs \<in> set_pur|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1116
      ==> \<exists>KC EKj.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1117
            Says C M {|{|sign (priSK C) MsgDualSign,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1118
                       EXcrypt KC EKj {|PIHead, Hash OIData|} PANData|},
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1119
                       OIData, Hash PIData|}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1120
              \<in> set evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1121
apply clarify
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 51798
diff changeset
  1122
apply hypsubst_thin
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1123
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1124
apply (erule rev_mp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1125
apply (erule set_pur.induct, simp_all, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1126
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1127
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1128
text{*When P receives an AuthReq and a dual signature, he knows that C and M
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1129
  agree on the essential details.  PurchAmt however is never sent by M to
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1130
  P; instead C and M both send 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1131
     @{term "HOD = Hash{|Number OrderDesc, Number PurchAmt|}"}
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1132
  and P compares the two copies of HOD.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1133
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1134
  Agreement can't be proved for some things, including the symmetric keys
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1135
  used in the digital envelopes.  On the other hand, M knows the true identity
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1136
  of PG (namely j'), and sends AReq there; he can't, however, check that
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1137
  the EXcrypt involves the correct PG's key.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1138
*}
15214
d3ab9b76ccb7 tidying up; identifying the main theorems
paulson
parents: 15205
diff changeset
  1139
theorem P_sees_CM_agreement:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1140
     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1141
         KC \<in> symKeys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1142
         Gets (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1143
           \<in> set evs;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1144
         C = Cardholder k;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1145
         PI_sign = sign (priSK C) {|Hash PIData, HOIData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1146
         P_I = {|PI_sign,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1147
                 EXcrypt KC (pubEK (PG j)) {|PIHead, HOIData|} PANData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1148
         PANData = {|Pan (pan C), Nonce (PANSecret k)|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1149
         PIData = {|PIHead, PANData|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1150
         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1151
                    TransStain|};
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1152
         evs \<in> set_pur;  C \<notin> bad;  M \<notin> bad|]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1153
  ==> \<exists>OIData OrderDesc KM' trans j' KC' KC'' P_I' P_I''.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1154
           HOD = Hash{|Number OrderDesc, Number PurchAmt|} &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1155
           HOIData = Hash OIData &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1156
           Notes M {|Number LID_M, Agent (PG j'), trans|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1157
           Says C M {|P_I', OIData, Hash PIData|} \<in> set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1158
           Says M (PG j') (EncB (priSK M) KM' (pubEK (PG j'))
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1159
                           AuthReqData P_I'')  \<in>  set evs &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1160
           P_I' = {|PI_sign,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1161
             EXcrypt KC' (pubEK (PG j')) {|PIHead, Hash OIData|} PANData|} &
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1162
           P_I'' = {|PI_sign,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1163
             EXcrypt KC'' (pubEK (PG j)) {|PIHead, Hash OIData|} PANData|}"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1164
apply clarify
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1165
apply (rule exE)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1166
apply (rule P_verifies_Signed_PReq [OF refl refl refl])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1167
apply (simp (no_asm_use) add: sign_def EncB_def, blast)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1168
apply (assumption+, clarify, simp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1169
apply (drule Gets_imp_knows_Spy [THEN parts.Inj], assumption)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1170
apply (blast elim: EncB_partsE dest: refl [THEN M_Says_AuthReq] unique_LID_M2)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1171
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1172
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
  1173
end