src/HOL/Auth/CertifiedEmail.thy
author haftmann
Fri, 11 Jun 2010 17:14:02 +0200
changeset 37407 61dd8c145da7
parent 35416 d8d7d1b785af
child 37936 1e4c5015a72e
permissions -rw-r--r--
declare lex_prod_def [code del]
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/CertifiedEmail
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
     2
    ID:         $Id$
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
     3
    Author:     Giampaolo Bella, Christiano Longo and Lawrence C Paulson
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
     4
*)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
     5
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
     6
header{*The Certified Electronic Mail Protocol by Abadi et al.*}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15068
diff changeset
     8
theory CertifiedEmail imports Public begin
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
     9
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 17689
diff changeset
    10
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    11
  TTP :: agent where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 17689
diff changeset
    12
  "TTP == Server"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    13
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    14
abbreviation
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    15
  RPwd :: "agent => key" where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 17689
diff changeset
    16
  "RPwd == shrK"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    17
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    18
 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    19
(*FIXME: the four options should be represented by pairs of 0 or 1.
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    20
  Right now only BothAuth is modelled.*)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    21
consts
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    22
  NoAuth   :: nat
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    23
  TTPAuth  :: nat
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    24
  SAuth    :: nat
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    25
  BothAuth :: nat
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    26
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    27
text{*We formalize a fixed way of computing responses.  Could be better.*}
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    28
definition "response" :: "agent => agent => nat => msg" where
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    29
   "response S R q == Hash {|Agent S, Key (shrK R), Nonce q|}"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    30
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    31
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    32
inductive_set certified_mail :: "event list set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    33
  where
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    34
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    35
  Nil: --{*The empty trace*}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    36
     "[] \<in> certified_mail"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    37
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    38
| Fake: --{*The Spy may say anything he can say.  The sender field is correct,
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    39
          but agents don't use that information.*}
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
    40
      "[| evsf \<in> certified_mail; X \<in> synth(analz(spies evsf))|] 
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
    41
       ==> Says Spy B X # evsf \<in> certified_mail"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    42
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    43
| FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    44
    equipped with the necessary credentials to serve as an SSL server.*}
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
    45
         "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    46
          ==> Notes TTP {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    47
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    48
| CM1: --{*The sender approaches the recipient.  The message is a number.*}
15068
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    49
 "[|evs1 \<in> certified_mail;
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    50
    Key K \<notin> used evs1;
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    51
    K \<in> symKeys;
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    52
    Nonce q \<notin> used evs1;
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    53
    hs = Hash{|Number cleartext, Nonce q, response S R q, Crypt K (Number m)|};
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    54
    S2TTP = Crypt(pubEK TTP) {|Agent S, Number BothAuth, Key K, Agent R, hs|}|]
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    55
  ==> Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
    56
                 Number cleartext, Nonce q, S2TTP|} # evs1 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
    57
        \<in> certified_mail"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    58
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    59
| CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    60
     password to @{term TTP} over an SSL channel.*}
15068
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    61
 "[|evs2 \<in> certified_mail;
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    62
    Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
    63
             Nonce q, S2TTP|} \<in> set evs2;
15068
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    64
    TTP \<noteq> R;  
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    65
    hr = Hash {|Number cleartext, Nonce q, response S R q, em|} |]
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    66
  ==> 
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    67
   Notes TTP {|Agent R, Agent TTP, S2TTP, Key(RPwd R), hr|} # evs2
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    68
      \<in> certified_mail"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    69
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    70
| CM3: --{*@{term TTP} simultaneously reveals the key to the recipient and gives
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    71
         a receipt to the sender.  The SSL channel does not authenticate 
14735
41d9efe3b5b1 removal of prime characters
paulson
parents: 14207
diff changeset
    72
         the client (@{term R}), but @{term TTP} accepts the message only 
41d9efe3b5b1 removal of prime characters
paulson
parents: 14207
diff changeset
    73
         if the given password is that of the claimed sender, @{term R}.
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    74
         He replies over the established SSL channel.*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    75
 "[|evs3 \<in> certified_mail;
14735
41d9efe3b5b1 removal of prime characters
paulson
parents: 14207
diff changeset
    76
    Notes TTP {|Agent R, Agent TTP, S2TTP, Key(RPwd R), hr|} \<in> set evs3;
41d9efe3b5b1 removal of prime characters
paulson
parents: 14207
diff changeset
    77
    S2TTP = Crypt (pubEK TTP) 
41d9efe3b5b1 removal of prime characters
paulson
parents: 14207
diff changeset
    78
                     {|Agent S, Number BothAuth, Key k, Agent R, hs|};
41d9efe3b5b1 removal of prime characters
paulson
parents: 14207
diff changeset
    79
    TTP \<noteq> R;  hs = hr;  k \<in> symKeys|]
15068
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    80
  ==> 
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    81
   Notes R {|Agent TTP, Agent R, Key k, hr|} # 
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    82
   Gets S (Crypt (priSK TTP) S2TTP) # 
58d216b32199 minor tweaks to go with the new version of the Accountability paper
paulson
parents: 14735
diff changeset
    83
   Says TTP S (Crypt (priSK TTP) S2TTP) # evs3 \<in> certified_mail"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    84
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    85
| Reception:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    86
 "[|evsr \<in> certified_mail; Says A B X \<in> set evsr|]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    87
  ==> Gets B X#evsr \<in> certified_mail"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    88
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    89
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
    90
declare Says_imp_knows_Spy [THEN analz.Inj, dest]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
    91
declare analz_into_parts [dest]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
    92
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    93
(*A "possibility property": there are traces that reach the end*)
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14145
diff changeset
    94
lemma "[| Key K \<notin> used []; K \<in> symKeys |] ==> 
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14145
diff changeset
    95
       \<exists>S2TTP. \<exists>evs \<in> certified_mail.
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
    96
           Says TTP S (Crypt (priSK TTP) S2TTP) \<in> set evs"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    97
apply (intro exI bexI)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    98
apply (rule_tac [2] certified_mail.Nil
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
    99
                    [THEN certified_mail.CM1, THEN certified_mail.Reception,
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   100
                     THEN certified_mail.CM2, 
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14145
diff changeset
   101
                     THEN certified_mail.CM3]) 
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14145
diff changeset
   102
apply (possibility, auto) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   103
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   104
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   105
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   106
lemma Gets_imp_Says:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   107
 "[| Gets B X \<in> set evs; evs \<in> certified_mail |] ==> \<exists>A. Says A B X \<in> set evs"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   108
apply (erule rev_mp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   109
apply (erule certified_mail.induct, auto)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   110
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   111
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   112
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   113
lemma Gets_imp_parts_knows_Spy:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   114
     "[|Gets A X \<in> set evs; evs \<in> certified_mail|] ==> X \<in> parts(spies evs)"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   115
apply (drule Gets_imp_Says, simp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   116
apply (blast dest: Says_imp_knows_Spy parts.Inj) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   117
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   118
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   119
lemma CM2_S2TTP_analz_knows_Spy:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   120
 "[|Gets R {|Agent A, Agent B, em, Number AO, Number cleartext, 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   121
              Nonce q, S2TTP|} \<in> set evs;
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   122
    evs \<in> certified_mail|] 
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   123
  ==> S2TTP \<in> analz(spies evs)"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   124
apply (drule Gets_imp_Says, simp) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   125
apply (blast dest: Says_imp_knows_Spy analz.Inj) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   126
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   127
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   128
lemmas CM2_S2TTP_parts_knows_Spy = 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   129
    CM2_S2TTP_analz_knows_Spy [THEN analz_subset_parts [THEN subsetD]]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   130
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   131
lemma hr_form_lemma [rule_format]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   132
 "evs \<in> certified_mail
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   133
  ==> hr \<notin> synth (analz (spies evs)) --> 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   134
      (\<forall>S2TTP. Notes TTP {|Agent R, Agent TTP, S2TTP, pwd, hr|}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   135
          \<in> set evs --> 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   136
      (\<exists>clt q S em. hr = Hash {|Number clt, Nonce q, response S R q, em|}))"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   137
apply (erule certified_mail.induct)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   138
apply (synth_analz_mono_contra, simp_all, blast+)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   139
done 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   140
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   141
text{*Cannot strengthen the first disjunct to @{term "R\<noteq>Spy"} because
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   142
the fakessl rule allows Spy to spoof the sender's name.  Maybe can
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   143
strengthen the second disjunct with @{term "R\<noteq>Spy"}.*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   144
lemma hr_form:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   145
 "[|Notes TTP {|Agent R, Agent TTP, S2TTP, pwd, hr|} \<in> set evs;
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   146
    evs \<in> certified_mail|]
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   147
  ==> hr \<in> synth (analz (spies evs)) | 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   148
      (\<exists>clt q S em. hr = Hash {|Number clt, Nonce q, response S R q, em|})"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   149
by (blast intro: hr_form_lemma) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   150
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   151
lemma Spy_dont_know_private_keys [dest!]:
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   152
    "[|Key (privateKey b A) \<in> parts (spies evs); evs \<in> certified_mail|]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   153
     ==> A \<in> bad"
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   154
apply (erule rev_mp) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   155
apply (erule certified_mail.induct, simp_all)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   156
txt{*Fake*}
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   157
apply (blast dest: Fake_parts_insert_in_Un) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   158
txt{*Message 1*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   159
apply blast  
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   160
txt{*Message 3*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   161
apply (frule_tac hr_form, assumption)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   162
apply (elim disjE exE) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   163
apply (simp_all add: parts_insert2) 
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   164
 apply (force dest!: parts_insert_subset_Un [THEN [2] rev_subsetD] 
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   165
                     analz_subset_parts [THEN subsetD], blast) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   166
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   167
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   168
lemma Spy_know_private_keys_iff [simp]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   169
    "evs \<in> certified_mail
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   170
     ==> (Key (privateKey b A) \<in> parts (spies evs)) = (A \<in> bad)"
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   171
by blast 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   172
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   173
lemma Spy_dont_know_TTPKey_parts [simp]:
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   174
     "evs \<in> certified_mail ==> Key (privateKey b TTP) \<notin> parts(spies evs)" 
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   175
by simp
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   176
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   177
lemma Spy_dont_know_TTPKey_analz [simp]:
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   178
     "evs \<in> certified_mail ==> Key (privateKey b TTP) \<notin> analz(spies evs)"
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   179
by auto
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   180
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   181
text{*Thus, prove any goal that assumes that @{term Spy} knows a private key
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   182
belonging to @{term TTP}*}
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   183
declare Spy_dont_know_TTPKey_parts [THEN [2] rev_notE, elim!]
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   184
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   185
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   186
lemma CM3_k_parts_knows_Spy:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   187
 "[| evs \<in> certified_mail;
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   188
     Notes TTP {|Agent A, Agent TTP,
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   189
                 Crypt (pubEK TTP) {|Agent S, Number AO, Key K, 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   190
                 Agent R, hs|}, Key (RPwd R), hs|} \<in> set evs|]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   191
  ==> Key K \<in> parts(spies evs)"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   192
apply (rotate_tac 1)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   193
apply (erule rev_mp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   194
apply (erule certified_mail.induct, simp_all)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   195
   apply (blast  intro:parts_insertI)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   196
txt{*Fake SSL*}
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   197
apply (blast dest: parts.Body) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   198
txt{*Message 2*}
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   199
apply (blast dest!: Gets_imp_Says elim!: knows_Spy_partsEs)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   200
txt{*Message 3*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   201
apply (frule_tac hr_form, assumption)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   202
apply (elim disjE exE) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   203
apply (simp_all add: parts_insert2)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   204
apply (blast intro: subsetD [OF parts_mono [OF Set.subset_insertI]])  
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   205
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   206
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   207
lemma Spy_dont_know_RPwd [rule_format]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   208
    "evs \<in> certified_mail ==> Key (RPwd A) \<in> parts(spies evs) --> A \<in> bad"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   209
apply (erule certified_mail.induct, simp_all) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   210
txt{*Fake*}
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   211
apply (blast dest: Fake_parts_insert_in_Un) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   212
txt{*Message 1*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   213
apply blast  
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   214
txt{*Message 3*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   215
apply (frule CM3_k_parts_knows_Spy, assumption)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   216
apply (frule_tac hr_form, assumption)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   217
apply (elim disjE exE) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   218
apply (simp_all add: parts_insert2) 
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   219
apply (force dest!: parts_insert_subset_Un [THEN [2] rev_subsetD]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   220
                    analz_subset_parts [THEN subsetD])
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   221
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   222
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   223
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   224
lemma Spy_know_RPwd_iff [simp]:
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   225
    "evs \<in> certified_mail ==> (Key (RPwd A) \<in> parts(spies evs)) = (A\<in>bad)"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   226
by (auto simp add: Spy_dont_know_RPwd) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   227
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   228
lemma Spy_analz_RPwd_iff [simp]:
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   229
    "evs \<in> certified_mail ==> (Key (RPwd A) \<in> analz(spies evs)) = (A\<in>bad)"
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   230
by (auto simp add: Spy_dont_know_RPwd [OF _ analz_subset_parts [THEN subsetD]])
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   231
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   232
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   233
text{*Unused, but a guarantee of sorts*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   234
theorem CertAutenticity:
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   235
     "[|Crypt (priSK TTP) X \<in> parts (spies evs); evs \<in> certified_mail|] 
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   236
      ==> \<exists>A. Says TTP A (Crypt (priSK TTP) X) \<in> set evs"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   237
apply (erule rev_mp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   238
apply (erule certified_mail.induct, simp_all) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   239
txt{*Fake*}
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   240
apply (blast dest: Spy_dont_know_private_keys Fake_parts_insert_in_Un)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   241
txt{*Message 1*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   242
apply blast 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   243
txt{*Message 3*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   244
apply (frule_tac hr_form, assumption)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   245
apply (elim disjE exE) 
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   246
apply (simp_all add: parts_insert2 parts_insert_knows_A) 
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   247
 apply (blast dest!: Fake_parts_sing_imp_Un, blast)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   248
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   249
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   250
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   251
subsection{*Proving Confidentiality Results*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   252
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   253
lemma analz_image_freshK [rule_format]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   254
 "evs \<in> certified_mail ==>
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   255
   \<forall>K KK. invKey (pubEK TTP) \<notin> KK -->
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   256
          (Key K \<in> analz (Key`KK Un (spies evs))) =
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   257
          (K \<in> KK | Key K \<in> analz (spies evs))"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   258
apply (erule certified_mail.induct)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   259
apply (drule_tac [6] A=TTP in symKey_neq_priEK) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   260
apply (erule_tac [6] disjE [OF hr_form]) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   261
apply (drule_tac [5] CM2_S2TTP_analz_knows_Spy) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   262
prefer 9
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   263
apply (elim exE)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   264
apply (simp_all add: synth_analz_insert_eq
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   265
                     subset_trans [OF _ subset_insertI]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   266
                     subset_trans [OF _ Un_upper2] 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   267
                del: image_insert image_Un add: analz_image_freshK_simps)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   268
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   269
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   270
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   271
lemma analz_insert_freshK:
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   272
  "[| evs \<in> certified_mail;  KAB \<noteq> invKey (pubEK TTP) |] ==>
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   273
      (Key K \<in> analz (insert (Key KAB) (spies evs))) =
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   274
      (K = KAB | Key K \<in> analz (spies evs))"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   275
by (simp only: analz_image_freshK analz_image_freshK_simps)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   276
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   277
text{*@{term S2TTP} must have originated from a valid sender
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   278
    provided @{term K} is secure.  Proof is surprisingly hard.*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   279
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   280
lemma Notes_SSL_imp_used:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   281
     "[|Notes B {|Agent A, Agent B, X|} \<in> set evs|] ==> X \<in> used evs"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   282
by (blast dest!: Notes_imp_used)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   283
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   284
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   285
(*The weaker version, replacing "used evs" by "parts (spies evs)", 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   286
   isn't inductive: message 3 case can't be proved *)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   287
lemma S2TTP_sender_lemma [rule_format]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   288
 "evs \<in> certified_mail ==>
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   289
    Key K \<notin> analz (spies evs) -->
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   290
    (\<forall>AO. Crypt (pubEK TTP)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   291
           {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   292
    (\<exists>m ctxt q. 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   293
        hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   294
        Says S R
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   295
           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   296
             Number ctxt, Nonce q,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   297
             Crypt (pubEK TTP)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   298
              {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))" 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   299
apply (erule certified_mail.induct, analz_mono_contra)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   300
apply (drule_tac [5] CM2_S2TTP_parts_knows_Spy, simp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   301
apply (simp add: used_Nil Crypt_notin_initState, simp_all)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   302
txt{*Fake*}
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   303
apply (blast dest: Fake_parts_sing [THEN subsetD]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   304
             dest!: analz_subset_parts [THEN subsetD])  
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   305
txt{*Fake SSL*}
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   306
apply (blast dest: Fake_parts_sing [THEN subsetD]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   307
             dest: analz_subset_parts [THEN subsetD])  
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   308
txt{*Message 1*}
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   309
apply (clarsimp, blast)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   310
txt{*Message 2*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   311
apply (simp add: parts_insert2, clarify) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   312
apply (drule parts_cut, assumption, simp) 
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   313
apply (blast intro: usedI) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   314
txt{*Message 3*} 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   315
apply (blast dest: Notes_SSL_imp_used used_parts_subset_parts) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   316
done 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   317
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   318
lemma S2TTP_sender:
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   319
 "[|Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs;
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   320
    Key K \<notin> analz (spies evs);
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   321
    evs \<in> certified_mail|]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   322
  ==> \<exists>m ctxt q. 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   323
        hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   324
        Says S R
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   325
           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   326
             Number ctxt, Nonce q,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   327
             Crypt (pubEK TTP)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   328
              {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs" 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   329
by (blast intro: S2TTP_sender_lemma) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   330
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   331
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   332
text{*Nobody can have used non-existent keys!*}
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   333
lemma new_keys_not_used [simp]:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   334
    "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> certified_mail|]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   335
     ==> K \<notin> keysFor (parts (spies evs))"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   336
apply (erule rev_mp) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   337
apply (erule certified_mail.induct, simp_all) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   338
txt{*Fake*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   339
apply (force dest!: keysFor_parts_insert) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   340
txt{*Message 1*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   341
apply blast 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   342
txt{*Message 3*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   343
apply (frule CM3_k_parts_knows_Spy, assumption)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   344
apply (frule_tac hr_form, assumption) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   345
apply (force dest!: keysFor_parts_insert)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   346
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   347
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   348
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   349
text{*Less easy to prove @{term "m'=m"}.  Maybe needs a separate unicity
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   350
theorem for ciphertexts of the form @{term "Crypt K (Number m)"}, 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   351
where @{term K} is secure.*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   352
lemma Key_unique_lemma [rule_format]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   353
     "evs \<in> certified_mail ==>
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   354
       Key K \<notin> analz (spies evs) -->
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   355
       (\<forall>m cleartext q hs.
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   356
        Says S R
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   357
           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   358
             Number cleartext, Nonce q,
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   359
             Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|}|}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   360
          \<in> set evs -->
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   361
       (\<forall>m' cleartext' q' hs'.
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   362
       Says S' R'
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   363
           {|Agent S', Agent TTP, Crypt K (Number m'), Number AO',
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   364
             Number cleartext', Nonce q',
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   365
             Crypt (pubEK TTP) {|Agent S', Number AO', Key K, Agent R', hs'|}|}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   366
          \<in> set evs --> R' = R & S' = S & AO' = AO & hs' = hs))" 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   367
apply (erule certified_mail.induct, analz_mono_contra, simp_all)
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   368
 prefer 2
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   369
 txt{*Message 1*}
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   370
 apply (blast dest!: Says_imp_knows_Spy [THEN parts.Inj] new_keys_not_used Crypt_imp_keysFor)
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   371
txt{*Fake*}
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   372
apply (auto dest!: usedI S2TTP_sender analz_subset_parts [THEN subsetD]) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   373
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   374
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   375
text{*The key determines the sender, recipient and protocol options.*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   376
lemma Key_unique:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   377
      "[|Says S R
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   378
           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   379
             Number cleartext, Nonce q,
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   380
             Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|}|}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   381
          \<in> set evs;
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   382
         Says S' R'
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   383
           {|Agent S', Agent TTP, Crypt K (Number m'), Number AO',
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   384
             Number cleartext', Nonce q',
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   385
             Crypt (pubEK TTP) {|Agent S', Number AO', Key K, Agent R', hs'|}|}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   386
          \<in> set evs;
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   387
         Key K \<notin> analz (spies evs);
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   388
         evs \<in> certified_mail|]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   389
       ==> R' = R & S' = S & AO' = AO & hs' = hs"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   390
by (rule Key_unique_lemma, assumption+)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   391
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   392
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   393
subsection{*The Guarantees for Sender and Recipient*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   394
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   395
text{*A Sender's guarantee:
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   396
      If Spy gets the key then @{term R} is bad and @{term S} moreover
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   397
      gets his return receipt (and therefore has no grounds for complaint).*}
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   398
theorem S_fairness_bad_R:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   399
      "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   400
                     Number cleartext, Nonce q, S2TTP|} \<in> set evs;
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   401
         S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   402
         Key K \<in> analz (spies evs);
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   403
         evs \<in> certified_mail;
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   404
         S\<noteq>Spy|]
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   405
      ==> R \<in> bad & Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   406
apply (erule rev_mp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   407
apply (erule ssubst)
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   408
apply (erule rev_mp)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   409
apply (erule certified_mail.induct, simp_all)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   410
txt{*Fake*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   411
apply spy_analz
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   412
txt{*Fake SSL*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   413
apply spy_analz
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   414
txt{*Message 3*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   415
apply (frule_tac hr_form, assumption)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   416
apply (elim disjE exE) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   417
apply (simp_all add: synth_analz_insert_eq  
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   418
                     subset_trans [OF _ subset_insertI]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   419
                     subset_trans [OF _ Un_upper2] 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   420
                del: image_insert image_Un add: analz_image_freshK_simps) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   421
apply (simp_all add: symKey_neq_priEK analz_insert_freshK)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   422
apply (blast dest: Notes_SSL_imp_used S2TTP_sender Key_unique)+
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   423
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   424
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   425
text{*Confidentially for the symmetric key*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   426
theorem Spy_not_see_encrypted_key:
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   427
      "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   428
                     Number cleartext, Nonce q, S2TTP|} \<in> set evs;
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   429
         S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   430
         evs \<in> certified_mail;
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   431
         S\<noteq>Spy; R \<notin> bad|]
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   432
      ==> Key K \<notin> analz(spies evs)"
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   433
by (blast dest: S_fairness_bad_R) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   434
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   435
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   436
text{*Agent @{term R}, who may be the Spy, doesn't receive the key
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   437
 until @{term S} has access to the return receipt.*} 
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   438
theorem S_guarantee:
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   439
     "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   440
                    Number cleartext, Nonce q, S2TTP|} \<in> set evs;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   441
        S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   442
        Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   443
        S\<noteq>Spy;  evs \<in> certified_mail|]
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   444
     ==> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   445
apply (erule rev_mp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   446
apply (erule ssubst)
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   447
apply (erule rev_mp)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   448
apply (erule certified_mail.induct, simp_all)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   449
txt{*Message 1*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   450
apply (blast dest: Notes_imp_used) 
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   451
txt{*Message 3*}
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   452
apply (blast dest: Notes_SSL_imp_used S2TTP_sender Key_unique S_fairness_bad_R) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   453
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   454
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   455
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   456
text{*If @{term R} sends message 2, and a delivery certificate exists, 
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   457
 then @{term R} receives the necessary key.  This result is also important
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   458
 to @{term S}, as it confirms the validity of the return receipt.*}
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16417
diff changeset
   459
theorem RR_validity:
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   460
  "[|Crypt (priSK TTP) S2TTP \<in> used evs;
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13934
diff changeset
   461
     S2TTP = Crypt (pubEK TTP)
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   462
               {|Agent S, Number AO, Key K, Agent R, 
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   463
                 Hash {|Number cleartext, Nonce q, r, em|}|};
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   464
     hr = Hash {|Number cleartext, Nonce q, r, em|};
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   465
     R\<noteq>Spy;  evs \<in> certified_mail|]
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   466
  ==> Notes R {|Agent TTP, Agent R, Key K, hr|} \<in> set evs"
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   467
apply (erule rev_mp)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   468
apply (erule ssubst)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   469
apply (erule ssubst)
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   470
apply (erule certified_mail.induct, simp_all)
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   471
txt{*Fake*} 
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   472
apply (blast dest: Fake_parts_sing [THEN subsetD]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   473
             dest!: analz_subset_parts [THEN subsetD])  
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   474
txt{*Fake SSL*}
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   475
apply (blast dest: Fake_parts_sing [THEN subsetD]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   476
            dest!: analz_subset_parts [THEN subsetD])  
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   477
txt{*Message 2*}
13934
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   478
apply (drule CM2_S2TTP_parts_knows_Spy, assumption)
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   479
apply (force dest: parts_cut)
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   480
txt{*Message 3*}
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   481
apply (frule_tac hr_form, assumption)
8c23dea4648e better guarantees
paulson
parents: 13926
diff changeset
   482
apply (elim disjE exE, simp_all) 
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   483
apply (blast dest: Fake_parts_sing [THEN subsetD]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 13956
diff changeset
   484
             dest!: analz_subset_parts [THEN subsetD]) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   485
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   486
75ae4244a596 Changes required by the certified email protocol
paulson
parents:
diff changeset
   487
end