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