src/HOL/Auth/Smartcard/ShoupRubinBella.thy
author desharna
Sat, 27 Nov 2021 10:46:57 +0100
changeset 74864 c256bba593f3
parent 69597 ff784d5a5bfb
child 80914 d97fdabd9e2b
permissions -rw-r--r--
redefined less_multiset to be based on multp
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32149
diff changeset
     1
(*  Author:     Giampaolo Bella, Catania University
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
     2
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
     3
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
     4
section\<open>Bella's modification of the Shoup-Rubin protocol\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
     5
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
     6
theory ShoupRubinBella imports Smartcard begin
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
     7
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
     8
text\<open>The modifications are that message 7 now mentions A, while message 10
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
     9
now mentions Nb and B. The lack of explicitness of the original version was
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    10
discovered by investigating adherence to the principle of Goal
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    11
Availability. Only the updated version makes the goals of confidentiality,
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    12
authentication and key distribution available to both peers.\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    13
41774
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    14
axiomatization sesK :: "nat*key => key"
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    15
where
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    16
   (*sesK is injective on each component*) 
41774
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    17
   inj_sesK [iff]: "(sesK(m,k) = sesK(m',k')) = (m = m' \<and> k = k')" and
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    18
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    19
   (*all long-term keys differ from sesK*)
41774
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    20
   shrK_disj_sesK [iff]: "shrK A \<noteq> sesK(m,pk)" and
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    21
   crdK_disj_sesK [iff]: "crdK C \<noteq> sesK(m,pk)" and
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    22
   pin_disj_sesK  [iff]: "pin P \<noteq> sesK(m,pk)" and
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    23
   pairK_disj_sesK[iff]: "pairK(A,B) \<noteq> sesK(m,pk)" and
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    24
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    25
   (*needed for base case in analz_image_freshK*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    26
   Atomic_distrib [iff]: "Atomic`(KEY`K \<union> NONCE`N) =
41774
13b97824aec6 modernized specifications;
wenzelm
parents: 35416
diff changeset
    27
                   Atomic`(KEY`K) \<union> Atomic`(NONCE`N)" and
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    28
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    29
  (*this protocol makes the assumption of secure means
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    30
    between each agent and his smartcard*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    31
   shouprubin_assumes_securemeans [iff]: "evs \<in> srb \<Longrightarrow> secureM"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    32
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32149
diff changeset
    33
definition Unique :: "[event, event list] => bool" ("Unique _ on _") where
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    34
   "Unique ev on evs == 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    35
      ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    36
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    37
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    38
inductive_set srb :: "event list set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    39
  where
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    40
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    41
    Nil:  "[]\<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    42
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    43
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    44
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    45
  | Fake: "\<lbrakk> evsF \<in> srb;  X \<in> synth (analz (knows Spy evsF)); 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    46
             illegalUse(Card B) \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    47
          \<Longrightarrow> Says Spy A X # 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    48
              Inputs Spy (Card B) X # evsF \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    49
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    50
(*In general this rule causes the assumption Card B \<notin> cloned
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    51
  in most guarantees for B - starting with confidentiality -
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    52
  otherwise pairK_confidential could not apply*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    53
  | Forge:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    54
         "\<lbrakk> evsFo \<in> srb; Nonce Nb \<in> analz (knows Spy evsFo);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    55
             Key (pairK(A,B)) \<in> knows Spy evsFo \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    56
          \<Longrightarrow> Notes Spy (Key (sesK(Nb,pairK(A,B)))) # evsFo \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    57
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    58
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    59
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    60
  | Reception: "\<lbrakk> evsrb\<in> srb; Says A B X \<in> set evsrb \<rbrakk>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    61
              \<Longrightarrow> Gets B X # evsrb \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    62
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    63
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    64
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    65
(*A AND THE SERVER*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    66
  | SR_U1:  "\<lbrakk> evs1 \<in> srb; A \<noteq> Server \<rbrakk>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    67
          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    68
                # evs1 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    69
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    70
  | SR_U2:  "\<lbrakk> evs2 \<in> srb; 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    71
             Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    72
          \<Longrightarrow> Says Server A \<lbrace>Nonce (Pairkey(A,B)), 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    73
                           Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    74
                  \<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    75
                # evs2 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    76
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    77
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    78
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    79
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    80
(*A AND HER CARD*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    81
(*A cannot decrypt the verifier for she dosn't know shrK A,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    82
  but the pairkey is recognisable*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    83
  | SR_U3:  "\<lbrakk> evs3 \<in> srb; legalUse(Card A);
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    84
             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    85
             Gets A \<lbrace>Nonce Pk, Certificate\<rbrace> \<in> set evs3 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    86
          \<Longrightarrow> Inputs A (Card A) (Agent A)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    87
                # evs3 \<in> srb"   (*however A only queries her card 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    88
if she has previously contacted the server to initiate with some B. 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    89
Otherwise she would do so even if the Server had not been active. 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    90
Still, this doesn't and can't mean that the pairkey originated with 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    91
the server*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    92
 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    93
(*The card outputs the nonce Na to A*)               
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
    94
  | SR_U4:  "\<lbrakk> evs4 \<in> srb; 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    95
             Nonce Na \<notin> used evs4; legalUse(Card A); A \<noteq> Server;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    96
             Inputs A (Card A) (Agent A) \<in> set evs4 \<rbrakk> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    97
       \<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    98
              # evs4 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
    99
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   100
(*The card can be exploited by the spy*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   101
(*because of the assumptions on the card, A is certainly not server nor spy*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   102
  | SR_U4Fake: "\<lbrakk> evs4F \<in> srb; Nonce Na \<notin> used evs4F; 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   103
             illegalUse(Card A);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   104
             Inputs Spy (Card A) (Agent A) \<in> set evs4F \<rbrakk> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   105
      \<Longrightarrow> Outpts (Card A) Spy \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   106
            # evs4F \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   107
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   108
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   109
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   110
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   111
(*A TOWARDS B*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   112
  | SR_U5:  "\<lbrakk> evs5 \<in> srb; 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   113
             Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs5;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   114
             \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   115
          \<Longrightarrow> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> # evs5 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   116
(*A must check that the verifier is not a compound message, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   117
  otherwise this would also fire after SR_U7 *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   118
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   119
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   120
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   121
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   122
(*B AND HIS CARD*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   123
  | SR_U6:  "\<lbrakk> evs6 \<in> srb; legalUse(Card B);
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   124
             Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs6 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   125
          \<Longrightarrow> Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   126
                # evs6 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   127
(*B gets back from the card the session key and various verifiers*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   128
  | SR_U7:  "\<lbrakk> evs7 \<in> srb; 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   129
             Nonce Nb \<notin> used evs7; legalUse(Card B); B \<noteq> Server;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   130
             K = sesK(Nb,pairK(A,B));
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   131
             Key K \<notin> used evs7;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   132
             Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7\<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   133
    \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   134
                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   135
                            Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   136
                # evs7 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   137
(*The card can be exploited by the spy*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   138
(*because of the assumptions on the card, A is certainly not server nor spy*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   139
  | SR_U7Fake:  "\<lbrakk> evs7F \<in> srb; Nonce Nb \<notin> used evs7F; 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   140
             illegalUse(Card B);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   141
             K = sesK(Nb,pairK(A,B));
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   142
             Key K \<notin> used evs7F;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   143
             Inputs Spy (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7F \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   144
          \<Longrightarrow> Outpts (Card B) Spy \<lbrace>Nonce Nb, Agent A, Key K,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   145
                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   146
                            Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   147
                # evs7F \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   148
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   149
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   150
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   151
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   152
(*B TOWARDS A*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   153
(*having sent an input that mentions A is the only memory B relies on,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   154
  since the output doesn't mention A - lack of explicitness*) 
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   155
  | SR_U8:  "\<lbrakk> evs8 \<in> srb;  
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   156
             Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs8;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   157
             Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   158
                                 Cert1, Cert2\<rbrace> \<in> set evs8 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   159
          \<Longrightarrow> Says B A \<lbrace>Nonce Nb, Cert1\<rbrace> # evs8 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   160
  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   161
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   162
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   163
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   164
(*A AND HER CARD*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   165
(*A cannot check the form of the verifiers - although I can prove the form of
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   166
  Cert2 - and just feeds her card with what she's got*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   167
  | SR_U9:  "\<lbrakk> evs9 \<in> srb; legalUse(Card A);
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   168
             Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs9;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   169
             Outpts (Card A) A \<lbrace>Nonce Na, Cert2\<rbrace> \<in> set evs9; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   170
             Gets A \<lbrace>Nonce Nb, Cert3\<rbrace> \<in> set evs9;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   171
             \<forall> p q. Cert2 \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   172
          \<Longrightarrow> Inputs A (Card A) 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   173
                 \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   174
                  Cert1, Cert3, Cert2\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   175
                # evs9 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   176
(*But the card will only give outputs to the inputs of the correct form*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   177
  | SR_U10: "\<lbrakk> evs10 \<in> srb; legalUse(Card A); A \<noteq> Server;
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   178
             K = sesK(Nb,pairK(A,B));
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   179
             Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   180
                                 Nonce (Pairkey(A,B)),
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   181
                                 Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   182
                                                   Agent B\<rbrace>,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   183
                                 Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   184
                                 Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   185
               \<in> set evs10 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   186
          \<Longrightarrow> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   187
                                 Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   188
                 # evs10 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   189
(*The card can be exploited by the spy*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   190
(*because of the assumptions on the card, A is certainly not server nor spy*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   191
  | SR_U10Fake: "\<lbrakk> evs10F \<in> srb; 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   192
             illegalUse(Card A);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   193
             K = sesK(Nb,pairK(A,B));
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   194
             Inputs Spy (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   195
                                   Nonce (Pairkey(A,B)),
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   196
                                   Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   197
                                                    Agent B\<rbrace>,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   198
                                   Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   199
                                   Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   200
               \<in> set evs10F \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   201
          \<Longrightarrow> Outpts (Card A) Spy \<lbrace>Agent B, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   202
                                   Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   203
                 # evs10F \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   204
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   205
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   206
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   207
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   208
(*A TOWARDS B*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   209
(*having initiated with B is the only memory A relies on,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   210
  since the output doesn't mention B - lack of explicitness*) 
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   211
  | SR_U11: "\<lbrakk> evs11 \<in> srb;
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   212
             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs11;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   213
             Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   214
               \<in> set evs11 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   215
          \<Longrightarrow> Says A B (Certificate) 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   216
                 # evs11 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   217
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   218
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   219
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   220
(*Both peers may leak by accident the session keys obtained from their
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   221
  cards*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   222
  | Oops1:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   223
     "\<lbrakk> evsO1 \<in> srb;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   224
         Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   225
           \<in> set evsO1 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   226
     \<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO1 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   227
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21588
diff changeset
   228
  | Oops2:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   229
     "\<lbrakk> evsO2 \<in> srb;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   230
         Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   231
           \<in> set evsO2 \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   232
    \<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO2 \<in> srb"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   233
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   234
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   235
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   236
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   237
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   238
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   239
(*To solve Fake case when it doesn't involve analz - used to be condensed
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   240
  into Fake_parts_insert_tac*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   241
declare Fake_parts_insert_in_Un  [dest]
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   242
declare analz_into_parts [dest]
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   243
(*declare parts_insertI [intro]*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   244
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   245
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   246
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   247
(*General facts about message reception*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   248
lemma Gets_imp_Says: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   249
       "\<lbrakk> Gets B X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> \<exists> A. Says A B X \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   250
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   251
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   252
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   253
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   254
lemma Gets_imp_knows_Spy: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   255
     "\<lbrakk> Gets B X \<in> set evs; evs \<in> srb \<rbrakk>  \<Longrightarrow> X \<in> knows Spy evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   256
apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   257
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   258
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   259
lemma Gets_imp_knows_Spy_parts_Snd: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   260
     "\<lbrakk> Gets B \<lbrace>X, Y\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  \<Longrightarrow> Y \<in> parts (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   261
apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy parts.Inj parts.Snd)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   262
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   263
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   264
lemma Gets_imp_knows_Spy_analz_Snd: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   265
     "\<lbrakk> Gets B \<lbrace>X, Y\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  \<Longrightarrow> Y \<in> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   266
apply (blast dest!: Gets_imp_Says Says_imp_knows_Spy analz.Inj analz.Snd)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   267
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   268
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   269
(*end general facts*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   270
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   271
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   272
24122
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 23894
diff changeset
   273
(*Begin lemmas on secure means, from Event.thy, proved for shouprubin. They help
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   274
  the simplifier, especially in analz_image_freshK*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   275
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   276
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   277
lemma Inputs_imp_knows_Spy_secureM_srb: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   278
      "\<lbrakk> Inputs Spy C X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> X \<in> knows Spy evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   279
apply (simp (no_asm_simp) add: Inputs_imp_knows_Spy_secureM)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   280
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   281
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   282
lemma knows_Spy_Inputs_secureM_srb_Spy: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   283
      "evs \<in>srb \<Longrightarrow> knows Spy (Inputs Spy C X # evs) = insert X (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   284
apply (simp (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   285
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   286
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   287
lemma knows_Spy_Inputs_secureM_srb: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   288
    "\<lbrakk> A \<noteq> Spy; evs \<in>srb \<rbrakk> \<Longrightarrow> knows Spy (Inputs A C X # evs) =  knows Spy evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   289
apply (simp (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   290
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   291
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   292
lemma knows_Spy_Outpts_secureM_srb_Spy: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   293
      "evs \<in>srb \<Longrightarrow> knows Spy (Outpts C Spy X # evs) = insert X (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   294
apply (simp (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   295
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   296
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   297
lemma knows_Spy_Outpts_secureM_srb: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   298
     "\<lbrakk> A \<noteq> Spy; evs \<in>srb \<rbrakk> \<Longrightarrow> knows Spy (Outpts C A X # evs) =  knows Spy evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   299
apply (simp (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   300
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   301
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   302
(*End lemmas on secure means for shouprubin*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   303
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   304
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   305
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   306
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   307
(*BEGIN technical lemmas - evolution of forwarding lemmas*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   308
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   309
(*If an honest agent uses a smart card, then the card is his/her own, is
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   310
  not stolen, and the agent has received suitable data to feed the card. 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   311
  In other words, these are guarantees that an honest agent can only use 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   312
  his/her own card, and must use it correctly.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   313
  On the contrary, the spy can "Inputs" any cloned cards also by the Fake rule.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   314
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   315
  Instead of Auto_tac, proofs here used to asm-simplify and then force-tac.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   316
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   317
lemma Inputs_A_Card_3: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   318
    "\<lbrakk> Inputs A C (Agent A) \<in> set evs; A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   319
     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   320
      (\<exists> Pk Certificate. Gets A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   321
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   322
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   323
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   324
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   325
lemma Inputs_B_Card_6: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   326
     "\<lbrakk> Inputs B C \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs; B \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   327
      \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and> Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   328
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   329
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   330
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   331
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   332
lemma Inputs_A_Card_9: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   333
     "\<lbrakk> Inputs A C \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   334
                                           Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   335
         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   336
  \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   337
      Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs     \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   338
      Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace> \<in> set evs        \<and>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   339
      Gets A \<lbrace>Nonce Nb, Cert2\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   340
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   341
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   342
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   343
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   344
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   345
(*The two occurrences of A in the Outpts event don't match SR_U4Fake, where
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   346
  A cannot be the Spy. Hence the card is legally usable by rule SR_U4*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   347
lemma Outpts_A_Card_4: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   348
     "\<lbrakk> Outpts C A \<lbrace>Nonce Na, (Crypt (crdK (Card A)) (Nonce Na))\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   349
         evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   350
     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   351
         Inputs A (Card A) (Agent A) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   352
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   353
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   354
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   355
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   356
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   357
(*First certificate is made explicit so that a comment similar to the previous
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   358
  applies. This also provides Na to the Inputs event in the conclusion*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   359
lemma Outpts_B_Card_7: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   360
      "\<lbrakk> Outpts C B \<lbrace>Nonce Nb, Agent A, Key K,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   361
                      Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   362
                      Cert2\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   363
         evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   364
     \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   365
         Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   366
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   367
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   368
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   369
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   370
lemma Outpts_A_Card_10: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   371
     "\<lbrakk> Outpts C A \<lbrace>Agent B, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   372
                    Key K, (Crypt (pairK(A,B)) (Nonce Nb))\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   373
         evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   374
     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   375
         (\<exists> Na Ver1 Ver2 Ver3.  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   376
       Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   377
                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   378
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   379
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   380
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   381
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   382
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   383
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   384
(*
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   385
Contrarily to original version, A doesn't need to check the form of the 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   386
certificate to learn that her peer is B. The goal is available to A.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   387
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   388
lemma Outpts_A_Card_10_imp_Inputs: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   389
     "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   390
          \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   391
     \<Longrightarrow> (\<exists> Na Ver1 Ver2 Ver3.  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   392
       Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   393
                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   394
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   395
apply simp_all
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   396
apply blast+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   397
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   398
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   399
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   400
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   401
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   402
(*Weaker version: if the agent can't check the forms of the verifiers, then
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   403
  the agent must not be the spy so as to solve SR_U4Fake. The verifier must be
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   404
  recognised as some cyphertex in order to distinguish from case SR_U7, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   405
  concerning B's output, which also begins with a nonce.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   406
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   407
lemma Outpts_honest_A_Card_4: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   408
     "\<lbrakk> Outpts C A \<lbrace>Nonce Na, Crypt K X\<rbrace> \<in>set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   409
         A \<noteq> Spy;  evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   410
     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   411
         Inputs A (Card A) (Agent A) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   412
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   413
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   414
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   415
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   416
(*alternative formulation of same theorem
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   417
Goal "\<lbrakk> Outpts C A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   418
         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>;    
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   419
         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   420
     \<Longrightarrow> legalUse(C) \<and> C = (Card A) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   421
         Inputs A (Card A) (Agent A) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   422
same proof
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   423
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   424
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   425
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   426
lemma Outpts_honest_B_Card_7: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   427
    "\<lbrakk> Outpts C B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   428
       B \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   429
   \<Longrightarrow> legalUse(C) \<and> C = (Card B) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   430
       (\<exists> Na. Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   431
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   432
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   433
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   434
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   435
lemma Outpts_honest_A_Card_10: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   436
     "\<lbrakk> Outpts C A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   437
         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   438
     \<Longrightarrow> legalUse (C) \<and> C = (Card A) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   439
         (\<exists> Na Pk Ver1 Ver2 Ver3.  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   440
          Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Pk,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   441
                              Ver1, Ver2, Ver3\<rbrace> \<in> set evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   442
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   443
apply simp_all
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   444
apply blast+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   445
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   446
(*-END-*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   447
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   448
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   449
(*Even weaker versions: if the agent can't check the forms of the verifiers
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   450
  and the agent may be the spy, then we must know what card the agent
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   451
  is getting the output from. 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   452
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   453
lemma Outpts_which_Card_4: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   454
    "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt K X\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   455
    \<Longrightarrow> Inputs A (Card A) (Agent A) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   456
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   457
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   458
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   459
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   460
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   461
lemma Outpts_which_Card_7: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   462
  "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   463
       \<in> set evs;  evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   464
     \<Longrightarrow> \<exists> Na. Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   465
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   466
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   467
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   468
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   469
(*This goal is now available - in the sense of Goal Availability*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   470
lemma Outpts_which_Card_10: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   471
    "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate \<rbrace> \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   472
       evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   473
    \<Longrightarrow> \<exists> Na. Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)), 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   474
                            Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   475
                            Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   476
                            Crypt (crdK (Card A)) (Nonce Na) \<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   477
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   478
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   479
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   480
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   481
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   482
(*Lemmas on the form of outputs*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   483
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   484
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   485
(*A needs to check that the verifier is a cipher for it to come from SR_U4
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   486
  otherwise it could come from SR_U7 *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   487
lemma Outpts_A_Card_form_4: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   488
  "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   489
         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   490
     \<Longrightarrow> Certificate = (Crypt (crdK (Card A)) (Nonce Na))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   491
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   492
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   493
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   494
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   495
lemma Outpts_B_Card_form_7: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   496
   "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   497
        \<in> set evs; evs \<in> srb \<rbrakk>          
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   498
      \<Longrightarrow> \<exists> Na.    
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   499
          K = sesK(Nb,pairK(A,B)) \<and>                       
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   500
          Cert1 = (Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   501
          Cert2 = (Crypt (pairK(A,B)) (Nonce Nb))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   502
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   503
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   504
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   505
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   506
lemma Outpts_A_Card_form_10: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   507
   "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   508
        \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   509
      \<Longrightarrow> K = sesK(Nb,pairK(A,B)) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   510
          Certificate = (Crypt (pairK(A,B)) (Nonce Nb))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   511
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   512
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   513
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   514
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   515
lemma Outpts_A_Card_form_bis: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   516
  "\<lbrakk> Outpts (Card A') A' \<lbrace>Agent B', Nonce Nb', Key (sesK(Nb,pairK(A,B))), 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   517
     Certificate\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   518
         evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   519
      \<Longrightarrow> A' = A \<and> B' = B \<and> Nb = Nb' \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   520
          Certificate = (Crypt (pairK(A,B)) (Nonce Nb))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   521
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   522
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   523
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   524
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   525
(*\<dots> and Inputs *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   526
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   527
lemma Inputs_A_Card_form_9: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   528
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   529
     "\<lbrakk> Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk,   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   530
                             Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   531
         evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   532
  \<Longrightarrow>    Cert3 = Crypt (crdK (Card A)) (Nonce Na)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   533
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   534
apply (erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   535
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   536
(*Fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   537
apply force
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   538
(*SR_U9*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   539
apply (blast dest!: Outpts_A_Card_form_4)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   540
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   541
(* Pk, Cert1, Cert2 cannot be made explicit because they traversed the network in the clear *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   542
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   543
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   544
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   545
(*General guarantees on Inputs and Outpts*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   546
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   547
(*for any agents*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   548
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   549
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   550
lemma Inputs_Card_legalUse: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   551
  "\<lbrakk> Inputs A (Card A) X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> legalUse(Card A)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   552
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   553
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   554
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   555
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   556
lemma Outpts_Card_legalUse: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   557
  "\<lbrakk> Outpts (Card A) A X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> legalUse(Card A)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   558
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   559
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   560
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   561
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   562
(*for honest agents*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   563
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   564
lemma Inputs_Card: "\<lbrakk> Inputs A C X \<in> set evs; A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   565
      \<Longrightarrow> C = (Card A) \<and> legalUse(C)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   566
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   567
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   568
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   569
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   570
lemma Outpts_Card: "\<lbrakk> Outpts C A X \<in> set evs; A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   571
      \<Longrightarrow> C = (Card A) \<and> legalUse(C)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   572
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   573
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   574
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   575
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   576
lemma Inputs_Outpts_Card: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   577
     "\<lbrakk> Inputs A C X \<in> set evs \<or> Outpts C A Y \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   578
         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   579
     \<Longrightarrow> C = (Card A) \<and> legalUse(Card A)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   580
apply (blast dest: Inputs_Card Outpts_Card)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   581
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   582
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   583
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   584
(*for the spy - they stress that the model behaves as it is meant to*) 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   585
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   586
(*The or version can be also proved directly.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   587
  It stresses that the spy may use either her own legally usable card or
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   588
  all the illegally usable cards.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   589
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   590
lemma Inputs_Card_Spy: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   591
  "\<lbrakk> Inputs Spy C X \<in> set evs \<or> Outpts C Spy X \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   592
      \<Longrightarrow> C = (Card Spy) \<and> legalUse(Card Spy) \<or>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   593
          (\<exists> A. C = (Card A) \<and> illegalUse(Card A))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   594
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   595
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   596
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   597
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   598
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   599
(*END technical lemmas*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   600
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   601
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   602
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   603
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   604
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   605
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   606
(*BEGIN unicity theorems: certain items uniquely identify a smart card's
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   607
                          output*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   608
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   609
(*A's card's first output: the nonce uniquely identifies the rest*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   610
lemma Outpts_A_Card_unique_nonce:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   611
     "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   612
           \<in> set evs;   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   613
         Outpts (Card A') A' \<lbrace>Nonce Na, Crypt (crdK (Card A')) (Nonce Na)\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   614
           \<in> set evs;   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   615
         evs \<in> srb \<rbrakk> \<Longrightarrow> A=A'"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   616
apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
   617
apply (fastforce dest: Outpts_parts_used)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   618
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   619
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   620
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   621
(*B's card's output: the NONCE uniquely identifies the rest*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   622
lemma Outpts_B_Card_unique_nonce: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   623
     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key SK, Cert1, Cert2\<rbrace> \<in> set evs;   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   624
      Outpts (Card B') B' \<lbrace>Nonce Nb, Agent A', Key SK', Cert1', Cert2'\<rbrace> \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   625
       evs \<in> srb \<rbrakk> \<Longrightarrow> B=B' \<and> A=A' \<and> SK=SK' \<and> Cert1=Cert1' \<and> Cert2=Cert2'"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   626
apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
   627
apply (fastforce dest: Outpts_parts_used)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   628
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   629
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   630
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   631
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   632
(*B's card's output: the SESKEY uniquely identifies the rest*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   633
lemma Outpts_B_Card_unique_key: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   634
     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key SK, Cert1, Cert2\<rbrace> \<in> set evs;   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   635
      Outpts (Card B') B' \<lbrace>Nonce Nb', Agent A', Key SK, Cert1', Cert2'\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   636
       evs \<in> srb \<rbrakk> \<Longrightarrow> B=B' \<and> A=A' \<and> Nb=Nb' \<and> Cert1=Cert1' \<and> Cert2=Cert2'"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   637
apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
   638
apply (fastforce dest: Outpts_parts_used)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   639
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   640
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   641
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   642
lemma Outpts_A_Card_unique_key: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   643
   "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, V\<rbrace> \<in> set evs;   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   644
      Outpts (Card A') A' \<lbrace>Agent B', Nonce Nb', Key K, V'\<rbrace> \<in> set evs;   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   645
         evs \<in> srb \<rbrakk> \<Longrightarrow> A=A' \<and> B=B' \<and> Nb=Nb' \<and> V=V'"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   646
apply (erule rev_mp, erule rev_mp, erule srb.induct, simp_all)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   647
apply (blast dest: Outpts_A_Card_form_bis)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   648
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   649
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   650
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   651
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   652
(*Revised unicity theorem - applies to both steps 4 and 7*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   653
lemma Outpts_A_Card_Unique: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   654
  "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, rest\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   655
     \<Longrightarrow> Unique (Outpts (Card A) A \<lbrace>Nonce Na, rest\<rbrace>) on evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   656
apply (erule rev_mp, erule srb.induct, simp_all add: Unique_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
   657
apply (fastforce dest: Outpts_parts_used)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   658
apply blast
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
   659
apply (fastforce dest: Outpts_parts_used)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   660
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   661
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   662
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   663
(*can't prove the same on evs10 for it doesn't have a freshness assumption!*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   664
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   665
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   666
(*END unicity theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   667
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   668
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   669
(*BEGIN counterguarantees about spy's knowledge*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   670
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   671
(*on nonces*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   672
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   673
lemma Spy_knows_Na: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   674
      "\<lbrakk> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   675
      \<Longrightarrow> Nonce Na \<in> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   676
apply (blast dest!: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   677
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   678
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   679
lemma Spy_knows_Nb: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   680
      "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   681
      \<Longrightarrow> Nonce Nb \<in> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   682
apply (blast dest!: Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Fst])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   683
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   684
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   685
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   686
(*on Pairkey*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   687
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   688
lemma Pairkey_Gets_analz_knows_Spy: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   689
      "\<lbrakk> Gets A \<lbrace>Nonce (Pairkey(A,B)), Certificate\<rbrace> \<in> set evs; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   690
      \<Longrightarrow> Nonce (Pairkey(A,B)) \<in> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   691
apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   692
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   693
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   694
lemma Pairkey_Inputs_imp_Gets: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   695
     "\<lbrakk> Inputs A (Card A)             
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   696
           \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),     
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   697
             Cert1, Cert3, Cert2\<rbrace> \<in> set evs;           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   698
         A \<noteq> Spy; evs \<in> srb \<rbrakk>     
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   699
      \<Longrightarrow> Gets A \<lbrace>Nonce (Pairkey(A,B)), Cert1\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   700
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   701
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   702
apply force
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   703
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   704
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   705
lemma Pairkey_Inputs_analz_knows_Spy: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   706
     "\<lbrakk> Inputs A (Card A)             
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   707
           \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce (Pairkey(A,B)),     
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   708
             Cert1, Cert3, Cert2\<rbrace> \<in> set evs;           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   709
         evs \<in> srb \<rbrakk>     
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   710
     \<Longrightarrow> Nonce (Pairkey(A,B)) \<in> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   711
apply (case_tac "A = Spy")
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
   712
apply (fastforce dest!: Inputs_imp_knows_Spy_secureM [THEN analz.Inj])
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   713
apply (blast dest!: Pairkey_Inputs_imp_Gets [THEN Pairkey_Gets_analz_knows_Spy])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   714
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   715
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   716
(* This fails on base case because of XOR properties.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   717
lemma Pairkey_authentic:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   718
  "\<lbrakk> Nonce (Pairkey(A,B)) \<in> parts (knows Spy evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   719
     Card A \<notin> cloned; evs \<in> sr \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   720
 \<Longrightarrow> \<exists> cert. Says Server A \<lbrace>Nonce (Pairkey(A,B)), Cert\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   721
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   722
apply (erule sr.induct, simp_all)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   723
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   724
oops
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   725
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   726
 1. \<And>x a b.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   727
       \<lbrakk>Card A \<notin> cloned; Pairkey (A, B) = Pairkey (a, b); Card a \<in> cloned;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   728
        Card b \<in> cloned\<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   729
       \<Longrightarrow> False
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   730
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   731
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   732
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   733
(*END counterguarantees on spy's knowledge*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   734
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   735
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   736
(*BEGIN rewrite rules for parts operator*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   737
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   738
declare shrK_disj_sesK [THEN not_sym, iff] 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   739
declare pin_disj_sesK [THEN not_sym, iff]
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   740
declare crdK_disj_sesK [THEN not_sym, iff]
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   741
declare pairK_disj_sesK [THEN not_sym, iff]
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   742
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   743
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   744
ML
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   745
\<open>
24122
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 23894
diff changeset
   746
structure ShoupRubinBella =
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 23894
diff changeset
   747
struct
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   748
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23746
diff changeset
   749
fun prepare_tac ctxt = 
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   750
 (*SR_U8*)   forward_tac ctxt [@{thm Outpts_B_Card_form_7}] 14 THEN
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42766
diff changeset
   751
 (*SR_U8*)   clarify_tac ctxt 15 THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   752
 (*SR_U9*)   forward_tac ctxt [@{thm Outpts_A_Card_form_4}] 16 THEN 
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   753
 (*SR_U11*)  forward_tac ctxt [@{thm Outpts_A_Card_form_10}] 21 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   754
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23746
diff changeset
   755
fun parts_prepare_tac ctxt = 
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23746
diff changeset
   756
           prepare_tac ctxt THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   757
 (*SR_U9*)   dresolve_tac ctxt [@{thm Gets_imp_knows_Spy_parts_Snd}] 18 THEN 
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   758
 (*SR_U9*)   dresolve_tac ctxt [@{thm Gets_imp_knows_Spy_parts_Snd}] 19 THEN 
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   759
 (*Oops1*) dresolve_tac ctxt [@{thm Outpts_B_Card_form_7}] 25    THEN               
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   760
 (*Oops2*) dresolve_tac ctxt [@{thm Outpts_A_Card_form_10}] 27 THEN                
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42766
diff changeset
   761
 (*Base*)  (force_tac ctxt) 1
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   762
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23746
diff changeset
   763
fun analz_prepare_tac ctxt = 
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23746
diff changeset
   764
         prepare_tac ctxt THEN
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   765
         dresolve_tac ctxt @{thms Gets_imp_knows_Spy_analz_Snd} 18 THEN 
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   766
 (*SR_U9*) dresolve_tac ctxt @{thms Gets_imp_knows_Spy_analz_Snd} 19 THEN 
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   767
         REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   768
24122
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 23894
diff changeset
   769
end
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   770
\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   771
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   772
method_setup prepare = \<open>
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   773
    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.prepare_tac ctxt))\<close>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 44890
diff changeset
   774
  "to launch a few simple facts that will help the simplifier"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   775
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   776
method_setup parts_prepare = \<open>
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   777
    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.parts_prepare_tac ctxt))\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   778
  "additional facts to reason about parts"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   779
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   780
method_setup analz_prepare = \<open>
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   781
    Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.analz_prepare_tac ctxt))\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   782
  "additional facts to reason about analz"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   783
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   784
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   785
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   786
lemma Spy_parts_keys [simp]: "evs \<in> srb \<Longrightarrow>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   787
  (Key (shrK P) \<in> parts (knows Spy evs)) = (Card P \<in> cloned) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   788
  (Key (pin P) \<in> parts (knows Spy evs)) = (P \<in> bad \<or> Card P \<in> cloned) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   789
  (Key (crdK C) \<in> parts (knows Spy evs)) = (C \<in> cloned) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   790
  (Key (pairK(A,B)) \<in> parts (knows Spy evs)) = (Card B \<in> cloned)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   791
apply (erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   792
apply parts_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   793
apply simp_all
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   794
apply (blast intro: parts_insertI)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   795
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   796
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   797
(*END rewrite rules for parts operator*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   798
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   799
(*BEGIN rewrite rules for analz operator*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   800
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   801
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   802
lemma Spy_analz_shrK[simp]: "evs \<in> srb \<Longrightarrow>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   803
  (Key (shrK P) \<in> analz (knows Spy evs)) = (Card P \<in> cloned)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   804
apply (auto dest!: Spy_knows_cloned)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   805
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   806
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   807
lemma Spy_analz_crdK[simp]: "evs \<in> srb \<Longrightarrow>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   808
  (Key (crdK C) \<in> analz (knows Spy evs)) = (C \<in> cloned)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   809
apply (auto dest!: Spy_knows_cloned)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   810
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   811
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   812
lemma Spy_analz_pairK[simp]: "evs \<in> srb \<Longrightarrow>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   813
  (Key (pairK(A,B)) \<in> analz (knows Spy evs)) = (Card B \<in> cloned)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   814
apply (auto dest!: Spy_knows_cloned)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   815
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   816
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   817
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   818
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   819
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   820
(*Because initState contains a set of nonces, this is needed for base case of
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   821
  analz_image_freshK*)
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61830
diff changeset
   822
lemma analz_image_Key_Un_Nonce:
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61830
diff changeset
   823
  "analz (Key ` K \<union> Nonce ` N) = Key ` K \<union> Nonce ` N"
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61830
diff changeset
   824
  by (auto simp del: parts_image)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   825
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   826
method_setup sc_analz_freshK = \<open>
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   827
    Scan.succeed (fn ctxt =>
30510
4120fc59dd85 unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents: 24122
diff changeset
   828
     (SIMPLE_METHOD
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   829
      (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   830
          REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshK_lemma}),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
   831
          ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt
24122
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 23894
diff changeset
   832
              addsimps [@{thm knows_Spy_Inputs_secureM_srb_Spy},
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 23894
diff changeset
   833
                  @{thm knows_Spy_Outpts_secureM_srb_Spy},
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 23894
diff changeset
   834
                  @{thm shouprubin_assumes_securemeans},
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   835
                  @{thm analz_image_Key_Un_Nonce}]))])))\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   836
    "for proving the Session Key Compromise theorem for smartcard protocols"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   837
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   838
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   839
lemma analz_image_freshK [rule_format]: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   840
     "evs \<in> srb \<Longrightarrow>      \<forall> K KK.  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   841
          (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) =  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   842
          (K \<in> KK \<or> Key K \<in> analz (knows Spy evs))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   843
apply (erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   844
apply analz_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   845
apply sc_analz_freshK
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   846
apply spy_analz
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   847
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   848
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   849
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   850
lemma analz_insert_freshK: "evs \<in> srb \<Longrightarrow>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   851
         Key K \<in> analz (insert (Key K') (knows Spy evs)) =  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   852
         (K = K' \<or> Key K \<in> analz (knows Spy evs))"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   853
apply (simp only: analz_image_freshK_simps analz_image_freshK)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   854
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   855
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   856
(*END rewrite rules for analz operator*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   857
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   858
(*BEGIN authenticity theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   859
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   860
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   861
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   862
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   863
lemma Na_Nb_certificate_authentic: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   864
     "\<lbrakk> Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace> \<in> parts (knows Spy evs);  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   865
         \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   866
         evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   867
     \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb,pairK(A,B))),   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   868
                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   869
                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   870
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   871
apply parts_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   872
apply simp_all
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   873
(*Fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   874
apply spy_analz
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   875
(*SR_U7F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   876
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   877
(*SR_U8*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   878
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   879
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   880
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   881
lemma Nb_certificate_authentic:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   882
      "\<lbrakk> Crypt (pairK(A,B)) (Nonce Nb) \<in> parts (knows Spy evs);  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   883
         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   884
         evs \<in> srb \<rbrakk>    
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   885
     \<Longrightarrow> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key (sesK(Nb,pairK(A,B))),  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   886
                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   887
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   888
apply parts_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   889
apply (case_tac [17] "Aa = Spy")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   890
apply simp_all
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   891
(*Fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   892
apply spy_analz
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   893
(*SR_U7F, SR_U10F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   894
apply clarify+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   895
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   896
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   897
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   898
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   899
(*Discovering the very origin of the Nb certificate...*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   900
lemma Outpts_A_Card_imp_pairK_parts: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   901
     "\<lbrakk> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   902
                    Key K, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   903
        evs \<in> srb \<rbrakk>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   904
    \<Longrightarrow> \<exists> Na. Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace> \<in> parts (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   905
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   906
apply parts_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   907
apply simp_all
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   908
(*Fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   909
apply (blast dest: parts_insertI)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   910
(*SR_U7*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   911
apply force
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   912
(*SR_U7F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   913
apply force
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   914
(*SR_U8*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   915
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   916
(*SR_U10*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   917
apply (blast dest: Inputs_imp_knows_Spy_secureM_srb parts.Inj Inputs_A_Card_9 Gets_imp_knows_Spy elim: knows_Spy_partsEs)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   918
(*SR_U10F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   919
apply (blast dest: Inputs_imp_knows_Spy_secureM_srb [THEN parts.Inj] 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   920
                   Inputs_A_Card_9 Gets_imp_knows_Spy 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   921
             elim: knows_Spy_partsEs)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   922
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   923
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   924
               
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   925
lemma Nb_certificate_authentic_bis: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   926
     "\<lbrakk> Crypt (pairK(A,B)) (Nonce Nb) \<in> parts (knows Spy evs);  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   927
         B \<noteq> Spy; \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   928
         evs \<in> srb \<rbrakk>    
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   929
 \<Longrightarrow> \<exists> Na. Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb,pairK(A,B))),   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   930
                   Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   931
                   Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   932
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   933
apply parts_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   934
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   935
(*Fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   936
apply spy_analz
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   937
(*SR_U7*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   938
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   939
(*SR_U7F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   940
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   941
(*SR_U8*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   942
apply force
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   943
(*SR_U10*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   944
apply (blast dest: Na_Nb_certificate_authentic Inputs_imp_knows_Spy_secureM_srb [THEN parts.Inj] elim: knows_Spy_partsEs)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   945
(*SR_U10F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   946
apply (blast dest: Na_Nb_certificate_authentic Inputs_imp_knows_Spy_secureM_srb [THEN parts.Inj] elim: knows_Spy_partsEs)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   947
(*SR_U11*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   948
apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_imp_pairK_parts)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   949
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   950
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   951
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   952
lemma Pairkey_certificate_authentic: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   953
    "\<lbrakk> Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace> \<in> parts (knows Spy evs);    
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   954
         Card A \<notin> cloned; evs \<in> srb \<rbrakk>        
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   955
     \<Longrightarrow> Pk = Pairkey(A,B) \<and>              
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   956
         Says Server A \<lbrace>Nonce Pk,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   957
                        Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   958
           \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   959
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   960
apply parts_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   961
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   962
(*Fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   963
apply spy_analz
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   964
(*SR_U8*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   965
apply force
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   966
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   967
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   968
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   969
lemma sesK_authentic: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   970
     "\<lbrakk> Key (sesK(Nb,pairK(A,B))) \<in> parts (knows Spy evs);  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   971
         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   972
         evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   973
      \<Longrightarrow> Notes Spy \<lbrace>Key (sesK(Nb,pairK(A,B))), Nonce Nb, Agent A, Agent B\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   974
           \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   975
apply (erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   976
apply parts_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   977
apply (simp_all)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   978
(*fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   979
apply spy_analz
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   980
(*forge*)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
   981
apply (fastforce dest: analz.Inj)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   982
(*SR_U7: used B\<noteq>Spy*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   983
(*SR_U7F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   984
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   985
(*SR_U10: used A\<noteq>Spy*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   986
(*SR_U10F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   987
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   988
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   989
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   990
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   991
(*END authenticity theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   992
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   993
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   994
(*BEGIN confidentiality theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   995
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   996
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   997
lemma Confidentiality: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   998
     "\<lbrakk> Notes Spy \<lbrace>Key (sesK(Nb,pairK(A,B))), Nonce Nb, Agent A, Agent B\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
   999
           \<notin> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1000
        A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1001
        evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1002
      \<Longrightarrow> Key (sesK(Nb,pairK(A,B))) \<notin> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1003
apply (blast intro: sesK_authentic)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1004
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1005
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1006
lemma Confidentiality_B: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1007
     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1008
          \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1009
        Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> \<notin> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1010
        A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); Card B \<notin> cloned; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1011
        evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1012
      \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1013
apply (erule rev_mp, erule rev_mp, erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1014
apply analz_prepare
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1015
apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1016
(*Fake*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1017
apply spy_analz
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1018
(*Forge*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1019
apply (rotate_tac 7)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1020
apply (drule parts.Inj)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
  1021
apply (fastforce dest: Outpts_B_Card_form_7)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1022
(*SR_U7*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1023
apply (blast dest!: Outpts_B_Card_form_7)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1024
(*SR_U7F*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1025
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1026
apply (drule Outpts_parts_used)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1027
apply simp
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1028
(*faster than
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
  1029
  apply (fastforce dest: Outpts_parts_used)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1030
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1031
(*SR_U10*)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
  1032
apply (fastforce dest: Outpts_B_Card_form_7)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1033
(*SR_U10F - uses assumption Card A not cloned*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1034
apply clarify
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1035
apply (drule Outpts_B_Card_form_7, assumption)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1036
apply simp
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1037
(*Oops1*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1038
apply (blast dest!: Outpts_B_Card_form_7)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1039
(*Oops2*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1040
apply (blast dest!: Outpts_B_Card_form_7 Outpts_A_Card_form_10)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1041
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1042
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1043
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1044
(*END confidentiality theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1045
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1046
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1047
(*BEGIN authentication theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1048
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1049
lemma A_authenticates_B: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1050
     "\<lbrakk> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1051
        \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1052
        evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1053
 \<Longrightarrow> \<exists> Na. Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K,   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1054
                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1055
                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1056
apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_form_10 Outpts_A_Card_imp_pairK_parts)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1057
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1058
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1059
lemma A_authenticates_B_Gets: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1060
     "\<lbrakk> Gets A \<lbrace>Nonce Nb, Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1061
           \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1062
         \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1063
         evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1064
    \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb, pairK (A, B))),   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1065
                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1066
                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1067
apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd, THEN Na_Nb_certificate_authentic])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1068
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1069
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1070
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1071
lemma A_authenticates_B_bis: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1072
     "\<lbrakk> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, Key K, Cert2\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1073
        \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1074
        evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1075
 \<Longrightarrow> \<exists> Cert1. Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1076
                \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1077
apply (blast dest: Na_Nb_certificate_authentic Outpts_A_Card_form_10 Outpts_A_Card_imp_pairK_parts)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1078
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1079
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1080
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1081
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1082
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1083
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1084
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1085
lemma B_authenticates_A: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1086
     "\<lbrakk> Gets B (Crypt (pairK(A,B)) (Nonce Nb)) \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1087
         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1088
         evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1089
      \<Longrightarrow> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1090
         Key (sesK(Nb,pairK(A,B))), Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1091
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1092
apply (erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1093
apply (simp_all (no_asm_simp))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1094
apply (blast dest: Says_imp_knows_Spy [THEN parts.Inj] Nb_certificate_authentic)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1095
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1096
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1097
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1098
lemma B_authenticates_A_bis: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1099
     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1100
        Gets B (Cert2) \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1101
         B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1102
         evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1103
      \<Longrightarrow> Outpts (Card A) A  \<lbrace>Agent B, Nonce Nb, Key K, Cert2\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1104
apply (blast dest: Outpts_B_Card_form_7 B_authenticates_A)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1105
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1106
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1107
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1108
(*END authentication theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1109
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1110
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1111
lemma Confidentiality_A: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1112
      "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1113
                       Key K, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1114
         Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> \<notin> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1115
         A \<noteq> Spy; B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1116
         evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1117
     \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1118
apply (drule A_authenticates_B)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1119
prefer 3
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1120
apply (erule exE)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1121
apply (drule Confidentiality_B)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1122
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1123
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1124
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1125
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1126
lemma Outpts_imp_knows_agents_secureM_srb: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1127
   "\<lbrakk> Outpts (Card A) A X \<in> set evs; evs \<in> srb \<rbrakk> \<Longrightarrow> X \<in> knows A evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1128
apply (simp (no_asm_simp) add: Outpts_imp_knows_agents_secureM)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1129
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1130
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1131
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1132
(*BEGIN key distribution theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1133
lemma A_keydist_to_B: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1134
   "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1135
         \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1136
         evs \<in> srb \<rbrakk>           
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1137
     \<Longrightarrow> Key K \<in> analz (knows B evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1138
apply (drule A_authenticates_B)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1139
prefer 3
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1140
apply (erule exE)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1141
apply (rule Outpts_imp_knows_agents_secureM_srb [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1142
apply assumption+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1143
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1144
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1145
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1146
lemma B_keydist_to_A: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1147
"\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1148
   Gets B (Cert2) \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1149
   B \<noteq> Spy; \<not>illegalUse(Card A); \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1150
   evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1151
 \<Longrightarrow> Key K \<in> analz (knows A evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1152
apply (frule Outpts_B_Card_form_7)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1153
apply assumption apply simp
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1154
apply (frule B_authenticates_A)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1155
apply (rule_tac [5] Outpts_imp_knows_agents_secureM_srb [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1156
apply simp+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1157
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1158
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1159
(*END key distribution theorems*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1160
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1161
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1162
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1163
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1164
(*BEGIN further theorems about authenticity of verifiers - useful to cards,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1165
  and somewhat to agents *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1166
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1167
(*MSG11
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1168
If B receives the verifier of msg11, then the verifier originated with msg7.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1169
This is clearly not available to B: B can't check the form of the verifier because he doesn't know pairK(A,B)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1170
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1171
lemma Nb_certificate_authentic_B: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1172
     "\<lbrakk> Gets B (Crypt (pairK(A,B)) (Nonce Nb)) \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1173
        B \<noteq> Spy; \<not>illegalUse(Card B); 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1174
        evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1175
     \<Longrightarrow> \<exists> Na. 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1176
            Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb,pairK(A,B))),   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1177
                Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1178
                Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1179
apply (blast dest: Gets_imp_knows_Spy [THEN parts.Inj, THEN Nb_certificate_authentic_bis])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1180
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1181
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1182
(*MSG10
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1183
If A obtains the verifier of msg10, then the verifier originated with msg7:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1184
A_authenticates_B. It is useful to A, who can check the form of the 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1185
verifier by application of Outpts_A_Card_form_10.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1186
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1187
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1188
(*MSG9
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1189
The first verifier verifies the Pairkey to the card: since it's encrypted
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1190
under Ka, it must come from the server (if A's card is not cloned).
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1191
The second verifier verifies both nonces, since it's encrypted under the
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1192
pairK, it must originate with B's card  (if A and B's cards not cloned).
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1193
The third verifier verifies Na: since it's encrytped under the card's key,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1194
it originated with the card; so the card does not need to save Na
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1195
in the first place and do a comparison now: it just verifies Na through the
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1196
verifier. Three theorems related to these three statements.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1197
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1198
Recall that a card can check the form of the verifiers (can decrypt them),
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1199
while an agent in general cannot, if not provided with a suitable theorem.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1200
*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1201
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1202
(*Card A can't reckon the pairkey - we need to guarantee its integrity!*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1203
lemma Pairkey_certificate_authentic_A_Card: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1204
     "\<lbrakk> Inputs A (Card A)   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1205
             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1206
               Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1207
               Cert2, Cert3\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1208
         A \<noteq> Spy; Card A \<notin> cloned; evs \<in> srb \<rbrakk>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1209
     \<Longrightarrow> Pk = Pairkey(A,B) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1210
         Says Server A \<lbrace>Nonce (Pairkey(A,B)),  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1211
                  Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>\<rbrace>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1212
           \<in> set evs "
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1213
apply (blast dest: Inputs_A_Card_9 Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd] Pairkey_certificate_authentic)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1214
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1215
(*the second conjunct of the thesis might be regarded as a form of integrity 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1216
  in the sense of Neuman-Ts'o*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1217
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1218
lemma Na_Nb_certificate_authentic_A_Card: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1219
      "\<lbrakk> Inputs A (Card A)   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1220
             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1221
          Cert1, Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, Cert3\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1222
      A \<noteq> Spy; \<not>illegalUse(Card B); evs \<in> srb \<rbrakk> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1223
   \<Longrightarrow> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key (sesK(Nb, pairK (A, B))),    
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1224
                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1225
                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1226
           \<in> set evs "
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1227
apply (frule Inputs_A_Card_9)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1228
apply assumption+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1229
apply (blast dest: Inputs_A_Card_9 Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd, THEN Na_Nb_certificate_authentic])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1230
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1231
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1232
lemma Na_authentic_A_Card: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1233
     "\<lbrakk> Inputs A (Card A)   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1234
             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1235
                Cert1, Cert2, Cert3\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1236
         A \<noteq> Spy; evs \<in> srb \<rbrakk>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1237
     \<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1238
           \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1239
apply (blast dest: Inputs_A_Card_9)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1240
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1241
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1242
(* These three theorems for Card A can be put together trivially.
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1243
They are separated to highlight the different requirements on agents
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1244
and their cards.*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1245
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1246
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1247
lemma Inputs_A_Card_9_authentic: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1248
  "\<lbrakk> Inputs A (Card A)   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1249
             \<lbrace>Agent B, Nonce Na, Nonce Nb, Nonce Pk, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1250
               Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1251
               Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>, Cert3\<rbrace> \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1252
    A \<noteq> Spy; Card A \<notin> cloned; \<not>illegalUse(Card B); evs \<in> srb \<rbrakk>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1253
    \<Longrightarrow>  Says Server A \<lbrace>Nonce Pk, Crypt (shrK A) \<lbrace>Nonce Pk, Agent B\<rbrace>\<rbrace>   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1254
           \<in> set evs  \<and> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1255
       Outpts (Card B) B \<lbrace>Nonce Nb, Agent A,  Key (sesK(Nb, pairK (A, B))),    
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1256
                             Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1257
                             Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1258
           \<in> set evs  \<and> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1259
         Outpts (Card A) A \<lbrace>Nonce Na, Cert3\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1260
           \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1261
apply (blast dest: Inputs_A_Card_9 Na_Nb_certificate_authentic Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Snd] Pairkey_certificate_authentic)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1262
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1263
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1264
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1265
(*MSG8
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1266
Nothing to prove because the message is a cleartext that comes from the 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1267
network*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1268
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1269
(*Other messages: nothing to prove because the verifiers involved are new*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1270
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1271
(*END further theorems about authenticity of verifiers*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1272
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1273
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1274
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1275
(* BEGIN trivial guarantees on outputs for agents *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1276
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1277
(*MSG4*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1278
lemma SR_U4_imp: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1279
     "\<lbrakk> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1280
           \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1281
         A \<noteq> Spy; evs \<in> srb \<rbrakk>                 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1282
     \<Longrightarrow> \<exists> Pk V. Gets A \<lbrace>Pk, V\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1283
apply (blast dest: Outpts_A_Card_4 Inputs_A_Card_3)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1284
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1285
(*weak: could strengthen the model adding verifier for the Pairkey to msg3*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1286
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1287
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1288
(*MSG7*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1289
lemma SR_U7_imp: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1290
     "\<lbrakk> Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1291
                      Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1292
                      Cert2\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1293
         B \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1294
     \<Longrightarrow> Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1295
apply (blast dest: Outpts_B_Card_7 Inputs_B_Card_6)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1296
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1297
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1298
(*MSG10*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1299
lemma SR_U10_imp: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1300
     "\<lbrakk> Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1301
                           Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1302
         \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1303
        A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1304
     \<Longrightarrow> \<exists> Cert1 Cert2.  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1305
                   Gets A \<lbrace>Nonce (Pairkey (A, B)), Cert1\<rbrace> \<in> set evs \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1306
                   Gets A \<lbrace>Nonce Nb, Cert2\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1307
apply (blast dest: Outpts_A_Card_10 Inputs_A_Card_9)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1308
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1309
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1310
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1311
(*END trivial guarantees on outputs for agents*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1312
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1313
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1314
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1315
(*INTEGRITY*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1316
lemma Outpts_Server_not_evs: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1317
      "evs \<in> srb \<Longrightarrow> Outpts (Card Server) P X \<notin> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1318
apply (erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1319
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1320
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1321
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62343
diff changeset
  1322
text\<open>\<^term>\<open>step2_integrity\<close> also is a reliability theorem\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1323
lemma Says_Server_message_form: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1324
     "\<lbrakk> Says Server A \<lbrace>Pk, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1325
         evs \<in> srb \<rbrakk>                   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1326
     \<Longrightarrow> \<exists> B. Pk = Nonce (Pairkey(A,B)) \<and>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1327
         Certificate = Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1328
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1329
apply (erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1330
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1331
apply (blast dest!: Outpts_Server_not_evs)+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1332
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1333
(*cannot be made useful to A in form of a Gets event*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1334
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
  1335
text\<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62343
diff changeset
  1336
  step4integrity is \<^term>\<open>Outpts_A_Card_form_4\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1337
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62343
diff changeset
  1338
  step7integrity is \<^term>\<open>Outpts_B_Card_form_7\<close>
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
  1339
\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1340
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1341
lemma step8_integrity: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1342
     "\<lbrakk> Says B A \<lbrace>Nonce Nb, Certificate\<rbrace> \<in> set evs;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1343
         B \<noteq> Server; B \<noteq> Spy; evs \<in> srb \<rbrakk>                   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1344
     \<Longrightarrow> \<exists> Cert2 K.   
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1345
    Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Certificate, Cert2\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1346
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1347
apply (erule srb.induct)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42793
diff changeset
  1348
prefer 18 apply (fastforce dest: Outpts_A_Card_form_10)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1349
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1350
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1351
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1352
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62343
diff changeset
  1353
text\<open>step9integrity is \<^term>\<open>Inputs_A_Card_form_9\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62343
diff changeset
  1354
        step10integrity is \<^term>\<open>Outpts_A_Card_form_10\<close>.
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
  1355
\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1356
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1357
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1358
lemma step11_integrity: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1359
     "\<lbrakk> Says A B (Certificate) \<in> set evs; 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1360
         \<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace>;  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1361
         A \<noteq> Spy; evs \<in> srb \<rbrakk>  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1362
     \<Longrightarrow> \<exists> K Nb.  
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1363
      Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1364
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1365
apply (erule srb.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1366
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1367
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1368
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1369
end
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents:
diff changeset
  1370