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