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