src/Doc/Tutorial/Protocol/NS_Public.thy
author wenzelm
Mon, 16 Sep 2019 15:30:38 +0200
changeset 70711 91319c3d2841
parent 69597 ff784d5a5bfb
child 76987 4c275405faae
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49322
fbb320d02420 tuned headers;
wenzelm
parents: 48985
diff changeset
     1
(*  Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     2
    Copyright   1996  University of Cambridge
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     3
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     4
Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol.
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     5
Version incorporating Lowe's fix (inclusion of B's identity in round 2).
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
     6
*)(*<*)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
     7
theory NS_Public imports Public begin(*>*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     8
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
     9
section\<open>Modelling the Protocol \label{sec:modelling}\<close>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    10
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
    11
text_raw \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    12
\begin{figure}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    13
\begin{isabelle}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
    14
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    15
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    16
inductive_set ns_public :: "event list set"
23733
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    17
  where
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    18
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    19
   Nil:  "[] \<in> ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    20
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    21
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    22
 | Fake: "\<lbrakk>evsf \<in> ns_public;  X \<in> synth (analz (knows Spy evsf))\<rbrakk>
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    23
          \<Longrightarrow> Says Spy B X  # evsf \<in> ns_public"
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    24
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    25
23733
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    26
 | NS1:  "\<lbrakk>evs1 \<in> ns_public;  Nonce NA \<notin> used evs1\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    27
          \<Longrightarrow> Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    28
                 # evs1  \<in>  ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    29
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    30
23733
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    31
 | NS2:  "\<lbrakk>evs2 \<in> ns_public;  Nonce NB \<notin> used evs2;
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    32
           Says A' B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    33
          \<Longrightarrow> Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    34
                # evs2  \<in>  ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    35
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    36
23733
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    37
 | NS3:  "\<lbrakk>evs3 \<in> ns_public;
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    38
           Says A  B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    39
           Says B' A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    40
              \<in> set evs3\<rbrakk>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    41
          \<Longrightarrow> Says A B (Crypt (pubK B) (Nonce NB)) # evs3 \<in> ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    42
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
    43
text_raw \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    44
\end{isabelle}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    45
\caption{An Inductive Protocol Definition}\label{fig:ns_public}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    46
\end{figure}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
    47
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    48
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
    49
text \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    50
Let us formalize the Needham-Schroeder public-key protocol, as corrected by
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    51
Lowe:
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    52
\begin{alignat*%
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    53
}{2}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    54
  &1.&\quad  A\to B  &: \comp{Na,A}\sb{Kb} \\
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    55
  &2.&\quad  B\to A  &: \comp{Na,Nb,B}\sb{Ka} \\
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    56
  &3.&\quad  A\to B  &: \comp{Nb}\sb{Kb}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    57
\end{alignat*%
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    58
}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    59
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    60
Each protocol step is specified by a rule of an inductive definition.  An
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    61
event trace has type \<open>event list\<close>, so we declare the constant
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    62
\<open>ns_public\<close> to be a set of such traces.
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    63
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    64
Figure~\ref{fig:ns_public} presents the inductive definition.  The
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    65
\<open>Nil\<close> rule introduces the empty trace.  The \<open>Fake\<close> rule models the
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    66
adversary's sending a message built from components taken from past
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    67
traffic, expressed using the functions \<open>synth\<close> and
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    68
\<open>analz\<close>. 
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    69
The next three rules model how honest agents would perform the three
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    70
protocol steps.  
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    71
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    72
Here is a detailed explanation of rule \<open>NS2\<close>.
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    73
A trace containing an event of the form
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    74
@{term [display,indent=5] "Says A' B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>)"}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    75
may be extended by an event of the form
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    76
@{term [display,indent=5] "Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)"}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    77
where \<open>NB\<close> is a fresh nonce: \<^term>\<open>Nonce NB \<notin> used evs2\<close>.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    78
Writing the sender as \<open>A'\<close> indicates that \<open>B\<close> does not 
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    79
know who sent the message.  Calling the trace variable \<open>evs2\<close> rather
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    80
than simply \<open>evs\<close> helps us know where we are in a proof after many
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    81
case-splits: every subgoal mentioning \<open>evs2\<close> involves message~2 of the
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    82
protocol.
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    83
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    84
Benefits of this approach are simplicity and clarity.  The semantic model
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    85
is set theory, proofs are by induction and the translation from the informal
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    86
notation to the inductive rules is straightforward. 
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
    87
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    88
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
    89
section\<open>Proving Elementary Properties \label{sec:regularity}\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    90
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
    91
(*<*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    92
declare knows_Spy_partsEs [elim]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    93
declare analz_subset_parts [THEN subsetD, dest]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    94
declare Fake_parts_insert [THEN subsetD, dest]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    95
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    96
(*A "possibility property": there are traces that reach the end*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    97
lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    98
apply (intro exI bexI)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    99
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2,
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   100
                                   THEN ns_public.NS3])
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   101
by possibility
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   102
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   103
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   104
(**** Inductive proofs about ns_public ****)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   105
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   106
(** Theorems of the form X \<notin> parts (knows Spy evs) imply that NOBODY
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   107
    sends messages containing X! **)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   108
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   109
(*Spy never sees another agent's private key! (unless it's bad at start)*)
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   110
(*>*)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   111
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   112
text \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   113
Secrecy properties can be hard to prove.  The conclusion of a typical
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   114
secrecy theorem is 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   115
\<^term>\<open>X \<notin> analz (knows Spy evs)\<close>.  The difficulty arises from
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   116
having to reason about \<open>analz\<close>, or less formally, showing that the spy
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   117
can never learn~\<open>X\<close>.  Much easier is to prove that \<open>X\<close> can never
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   118
occur at all.  Such \emph{regularity} properties are typically expressed
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   119
using \<open>parts\<close> rather than \<open>analz\<close>.
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   120
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   121
The following lemma states that \<open>A\<close>'s private key is potentially
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   122
known to the spy if and only if \<open>A\<close> belongs to the set \<open>bad\<close> of
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   123
compromised agents.  The statement uses \<open>parts\<close>: the very presence of
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   124
\<open>A\<close>'s private key in a message, whether protected by encryption or
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   125
not, is enough to confirm that \<open>A\<close> is compromised.  The proof, like
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   126
nearly all protocol proofs, is by induction over traces.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   127
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   128
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   129
lemma Spy_see_priK [simp]:
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   130
      "evs \<in> ns_public
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   131
       \<Longrightarrow> (Key (priK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   132
apply (erule ns_public.induct, simp_all)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   133
txt \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   134
The induction yields five subgoals, one for each rule in the definition of
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   135
\<open>ns_public\<close>.  The idea is to prove that the protocol property holds initially
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   136
(rule \<open>Nil\<close>), is preserved by each of the legitimate protocol steps (rules
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   137
\<open>NS1\<close>--\<open>3\<close>), and even is preserved in the face of anything the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   138
spy can do (rule \<open>Fake\<close>).  
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   139
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   140
The proof is trivial.  No legitimate protocol rule sends any keys
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   141
at all, so only \<open>Fake\<close> is relevant. Indeed, simplification leaves
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   142
only the \<open>Fake\<close> case, as indicated by the variable name \<open>evsf\<close>:
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   143
@{subgoals[display,indent=0,margin=65]}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   144
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   145
by blast
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   146
(*<*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   147
lemma Spy_analz_priK [simp]:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   148
      "evs \<in> ns_public \<Longrightarrow> (Key (priK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   149
by auto
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   150
(*>*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   151
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   152
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   153
The \<open>Fake\<close> case is proved automatically.  If
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   154
\<^term>\<open>priK A\<close> is in the extended trace then either (1) it was already in the
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   155
original trace or (2) it was
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   156
generated by the spy, who must have known this key already. 
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   157
Either way, the induction hypothesis applies.
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   158
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   159
\emph{Unicity} lemmas are regularity lemmas stating that specified items
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   160
can occur only once in a trace.  The following lemma states that a nonce
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   161
cannot be used both as $Na$ and as $Nb$ unless
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   162
it is known to the spy.  Intuitively, it holds because honest agents
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   163
always choose fresh values as nonces; only the spy might reuse a value,
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   164
and he doesn't know this particular value.  The proof script is short:
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   165
induction, simplification, \<open>blast\<close>.  The first line uses the rule
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   166
\<open>rev_mp\<close> to prepare the induction by moving two assumptions into the 
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   167
induction formula.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   168
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   169
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   170
lemma no_nonce_NS1_NS2:
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   171
    "\<lbrakk>Crypt (pubK C) \<lbrace>NA', Nonce NA, Agent D\<rbrace> \<in> parts (knows Spy evs);
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   172
      Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs);
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   173
      evs \<in> ns_public\<rbrakk>
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   174
     \<Longrightarrow> Nonce NA \<in> analz (knows Spy evs)"
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   175
apply (erule rev_mp, erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   176
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   177
apply (blast intro: analz_insertI)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   178
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   179
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   180
text \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   181
The following unicity lemma states that, if \isa{NA} is secret, then its
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   182
appearance in any instance of message~1 determines the other components. 
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   183
The proof is similar to the previous one.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   184
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   185
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   186
lemma unique_NA:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   187
     "\<lbrakk>Crypt(pubK B)  \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   188
       Crypt(pubK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   189
       Nonce NA \<notin> analz (knows Spy evs); evs \<in> ns_public\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   190
      \<Longrightarrow> A=A' \<and> B=B'"
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   191
(*<*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   192
apply (erule rev_mp, erule rev_mp, erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   193
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   194
(*Fake, NS1*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   195
apply (blast intro: analz_insertI)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   196
done
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   197
(*>*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   198
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   199
section\<open>Proving Secrecy Theorems \label{sec:secrecy}\<close>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   200
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   201
(*<*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   202
(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   203
  The major premise "Says A B ..." makes it a dest-rule, so we use
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   204
  (erule rev_mp) rather than rule_format. *)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   205
theorem Spy_not_see_NA:
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   206
      "\<lbrakk>Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   207
        A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   208
       \<Longrightarrow> Nonce NA \<notin> analz (knows Spy evs)"
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   209
apply (erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   210
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   211
apply spy_analz
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   212
apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   213
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   214
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   215
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   216
(*Authentication for A: if she receives message 2 and has used NA
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   217
  to start a run, then B has sent message 2.*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   218
lemma A_trusts_NS2_lemma [rule_format]:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   219
   "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   220
    \<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32891
diff changeset
   221
        Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32891
diff changeset
   222
        Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   223
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   224
(*Fake, NS1*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   225
apply (blast dest: Spy_not_see_NA)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   226
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   227
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   228
theorem A_trusts_NS2:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   229
     "\<lbrakk>Says A  B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   230
       Says B' A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   231
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   232
      \<Longrightarrow> Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   233
by (blast intro: A_trusts_NS2_lemma)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   234
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   235
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   236
(*If the encrypted message appears then it originated with Alice in NS1*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   237
lemma B_trusts_NS1 [rule_format]:
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   238
     "evs \<in> ns_public
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   239
      \<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32891
diff changeset
   240
          Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32891
diff changeset
   241
          Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   242
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   243
(*Fake*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   244
apply (blast intro!: analz_insertI)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   245
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   246
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   247
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   248
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   249
(*** Authenticity properties obtained from NS2 ***)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   250
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   251
(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   252
  [unicity of B makes Lowe's fix work]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   253
  [proof closely follows that for unique_NA] *)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   254
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   255
lemma unique_NB [dest]:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   256
     "\<lbrakk>Crypt(pubK A)  \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   257
       Crypt(pubK A') \<lbrace>Nonce NA', Nonce NB, Agent B'\<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   258
       Nonce NB \<notin> analz (knows Spy evs); evs \<in> ns_public\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   259
      \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B'"
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   260
apply (erule rev_mp, erule rev_mp, erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   261
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   262
(*Fake, NS2*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   263
apply (blast intro: analz_insertI)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   264
done
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   265
(*>*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   266
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   267
text \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   268
The secrecy theorems for Bob (the second participant) are especially
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   269
important because they fail for the original protocol.  The following
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   270
theorem states that if Bob sends message~2 to Alice, and both agents are
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   271
uncompromised, then Bob's nonce will never reach the spy.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   272
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   273
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   274
theorem Spy_not_see_NB [dest]:
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   275
 "\<lbrakk>Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   276
   A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   277
  \<Longrightarrow> Nonce NB \<notin> analz (knows Spy evs)"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   278
txt \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   279
To prove it, we must formulate the induction properly (one of the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   280
assumptions mentions~\<open>evs\<close>), apply induction, and simplify:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   281
\<close>
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   282
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   283
apply (erule rev_mp, erule ns_public.induct, simp_all)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   284
(*<*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   285
apply spy_analz
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   286
defer
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   287
apply (blast intro: no_nonce_NS1_NS2)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   288
apply (blast intro: no_nonce_NS1_NS2)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   289
(*>*)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   290
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   291
txt \<open>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   292
The proof states are too complicated to present in full.  
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   293
Let's examine the simplest subgoal, that for message~1.  The following
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   294
event has just occurred:
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   295
\[ 1.\quad  A'\to B'  : \comp{Na',A'}\sb{Kb'} \]
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   296
The variables above have been primed because this step
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   297
belongs to a different run from that referred to in the theorem
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   298
statement --- the theorem
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   299
refers to a past instance of message~2, while this subgoal
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   300
concerns message~1 being sent just now.
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   301
In the Isabelle subgoal, instead of primed variables like $B'$ and $Na'$
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   302
we have \<open>Ba\<close> and~\<open>NAa\<close>:
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   303
@{subgoals[display,indent=0]}
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   304
The simplifier has used a 
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   305
default simplification rule that does a case
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   306
analysis for each encrypted message on whether or not the decryption key
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   307
is compromised.
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   308
@{named_thms [display,indent=0,margin=50] analz_Crypt_if [no_vars] (analz_Crypt_if)}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   309
The simplifier has also used \<open>Spy_see_priK\<close>, proved in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   310
{\S}\ref{sec:regularity} above, to yield \<^term>\<open>Ba \<in> bad\<close>.
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   311
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   312
Recall that this subgoal concerns the case
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   313
where the last message to be sent was
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   314
\[ 1.\quad  A'\to B'  : \comp{Na',A'}\sb{Kb'}. \]
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   315
This message can compromise $Nb$ only if $Nb=Na'$ and $B'$ is compromised,
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   316
allowing the spy to decrypt the message.  The Isabelle subgoal says
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   317
precisely this, if we allow for its choice of variable names.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   318
Proving \<^term>\<open>NB \<noteq> NAa\<close> is easy: \<open>NB\<close> was
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   319
sent earlier, while \<open>NAa\<close> is fresh; formally, we have
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   320
the assumption \<^term>\<open>Nonce NAa \<notin> used evs1\<close>. 
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   321
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   322
Note that our reasoning concerned \<open>B\<close>'s participation in another
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   323
run.  Agents may engage in several runs concurrently, and some attacks work
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   324
by interleaving the messages of two runs.  With model checking, this
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   325
possibility can cause a state-space explosion, and for us it
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   326
certainly complicates proofs.  The biggest subgoal concerns message~2.  It
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   327
splits into several cases, such as whether or not the message just sent is
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   328
the very message mentioned in the theorem statement.
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   329
Some of the cases are proved by unicity, others by
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   330
the induction hypothesis.  For all those complications, the proofs are
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   331
automatic by \<open>blast\<close> with the theorem \<open>no_nonce_NS1_NS2\<close>.
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   332
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   333
The remaining theorems about the protocol are not hard to prove.  The
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   334
following one asserts a form of \emph{authenticity}: if
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   335
\<open>B\<close> has sent an instance of message~2 to~\<open>A\<close> and has received the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   336
expected reply, then that reply really originated with~\<open>A\<close>.  The
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   337
proof is a simple induction.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   338
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   339
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   340
(*<*)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   341
by (blast intro: no_nonce_NS1_NS2)
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   342
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   343
lemma B_trusts_NS3_lemma [rule_format]:
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   344
     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   345
      Crypt (pubK B) (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   346
      Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   347
      Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   348
by (erule ns_public.induct, auto)
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   349
(*>*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   350
theorem B_trusts_NS3:
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   351
 "\<lbrakk>Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   352
   Says A' B (Crypt (pubK B) (Nonce NB)) \<in> set evs;
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   353
   A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   354
  \<Longrightarrow> Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   355
(*<*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   356
by (blast intro: B_trusts_NS3_lemma)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   357
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   358
(*** Overall guarantee for B ***)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   359
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   360
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   361
(*If NS3 has been sent and the nonce NB agrees with the nonce B joined with
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   362
  NA, then A initiated the run using NA.*)
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   363
theorem B_trusts_protocol [rule_format]:
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   364
     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   365
      Crypt (pubK B) (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   366
      Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   367
      Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   368
by (erule ns_public.induct, auto)
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   369
(*>*)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   370
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   371
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   372
From similar assumptions, we can prove that \<open>A\<close> started the protocol
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   373
run by sending an instance of message~1 involving the nonce~\<open>NA\<close>\@. 
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   374
For this theorem, the conclusion is 
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 27093
diff changeset
   375
@{thm [display] (concl) B_trusts_protocol [no_vars]}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   376
Analogous theorems can be proved for~\<open>A\<close>, stating that nonce~\<open>NA\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   377
remains secret and that message~2 really originates with~\<open>B\<close>.  Even the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
   378
flawed protocol establishes these properties for~\<open>A\<close>;
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   379
the flaw only harms the second participant.
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   380
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   381
\medskip
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   382
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   383
Detailed information on this protocol verification technique can be found
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 56739
diff changeset
   384
elsewhere~@{cite "paulson-jcs"}, including proofs of an Internet
7435b6a3f72e more antiquotations;
wenzelm
parents: 56739
diff changeset
   385
protocol~@{cite "paulson-tls"}.  We must stress that the protocol discussed
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   386
in this chapter is trivial.  There are only three messages; no keys are
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   387
exchanged; we merely have to prove that encrypted data remains secret. 
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   388
Real world protocols are much longer and distribute many secrets to their
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   389
participants.  To be realistic, the model has to include the possibility
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   390
of keys being lost dynamically due to carelessness.  If those keys have
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   391
been used to encrypt other sensitive information, there may be cascading
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   392
losses.  We may still be able to establish a bound on the losses and to
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   393
prove that other protocol runs function
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 56739
diff changeset
   394
correctly~@{cite "paulson-yahalom"}.  Proofs of real-world protocols follow
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   395
the strategy illustrated above, but the subgoals can
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   396
be much bigger and there are more of them.
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   397
\index{protocols!security|)}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58620
diff changeset
   398
\<close>
23925
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   399
ee98c2528a8f LaTeX code is now generated directly from theory files.
berghofe
parents: 23733
diff changeset
   400
(*<*)end(*>*)