src/HOL/Auth/NS_Public_Bad.thy
author Simon Wimmer <wimmers@in.tum.de>
Thu, 18 Apr 2024 17:53:14 +0200
changeset 80137 0c51e0a6bc37
parent 76368 943f99825f39
permissions -rw-r--r--
sketch & explore: recover from duplicate fixed variables in Isar proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 32960
diff changeset
     1
(*  Title:      HOL/Auth/NS_Public_Bad.thy
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     3
    Copyright   1996  University of Cambridge
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     4
*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     5
76299
0ad6f6508274 Tidying of some very old proofs
paulson <lp15@cam.ac.uk>
parents: 76297
diff changeset
     6
section\<open>The Needham-Schroeder Public-Key Protocol (Flawed)\<close>
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13922
diff changeset
     7
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
     8
text \<open>Flawed version, vulnerable to Lowe's attack.
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
     9
From Burrows, Abadi and Needham.  A Logic of Authentication.
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    10
  Proc. Royal Soc. 426 (1989), p. 260\<close>
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    11
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14200
diff changeset
    12
theory NS_Public_Bad imports Public begin
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    13
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    14
inductive_set ns_public :: "event list set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    15
  where
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    16
   Nil:  "[] \<in> ns_public" 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    17
   \<comment> \<open>Initial trace is empty\<close>
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    18
 | Fake: "\<lbrakk>evsf \<in> ns_public;  X \<in> synth (analz (spies evsf))\<rbrakk>
11366
b42287eb20cf renaming of evs in the Fake rule
paulson
parents: 11230
diff changeset
    19
          \<Longrightarrow> Says Spy B X  # evsf \<in> ns_public"
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    20
   \<comment> \<open>The spy can say almost anything.\<close>
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    21
 | NS1:  "\<lbrakk>evs1 \<in> ns_public;  Nonce NA \<notin> used evs1\<rbrakk>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    22
          \<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    23
                # evs1  \<in>  ns_public"
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    24
   \<comment> \<open>Alice initiates a protocol run, sending a nonce to Bob\<close>
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    25
 | NS2:  "\<lbrakk>evs2 \<in> ns_public;  Nonce NB \<notin> used evs2;
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    26
           Says A' B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    27
          \<Longrightarrow> Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    28
                # evs2  \<in>  ns_public"
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    29
   \<comment> \<open>Bob responds to Alice's message with a further nonce\<close>
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    30
 | NS3:  "\<lbrakk>evs3 \<in> ns_public;
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    31
           Says A  B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    32
           Says B' A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs3\<rbrakk>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    33
          \<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) # evs3 \<in> ns_public"
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    34
   \<comment> \<open>Alice proves her existence by sending @{term NB} back to Bob.\<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    35
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    36
declare knows_Spy_partsEs [elim]
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
    37
declare analz_into_parts [dest]
56681
e8d5d60d655e avoid non-standard simp default rule
haftmann
parents: 38964
diff changeset
    38
declare Fake_parts_insert_in_Un [dest]
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    39
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    40
text \<open>A "possibility property": there are traces that reach the end\<close>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    41
lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    42
  apply (intro exI bexI)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    43
   apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2, THEN ns_public.NS3])
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    44
  by possibility
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    45
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    46
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    47
subsection \<open>Inductive proofs about @{term ns_public}\<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    48
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    49
(** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    50
    sends messages containing X! **)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    51
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    52
text \<open>Spy never sees another agent's private key! (unless it's bad at start)\<close>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    53
lemma Spy_see_priEK [simp]: 
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    54
  "evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> parts (spies evs)) = (A \<in> bad)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    55
  by (erule ns_public.induct, auto)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    56
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    57
lemma Spy_analz_priEK [simp]: 
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    58
  "evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> analz (spies evs)) = (A \<in> bad)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    59
  by auto
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    60
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    61
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    62
subsection \<open>Authenticity properties obtained from {term NS1}\<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    63
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    64
text \<open>It is impossible to re-use a nonce in both {term NS1} and {term NS2}, provided the nonce
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    65
  is secret.  (Honest users generate fresh nonces.)\<close>
76368
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
    66
lemma no_nonce_NS1_NS2: 
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
    67
      "\<lbrakk>evs \<in> ns_public;
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
    68
        Crypt (pubEK C) \<lbrace>NA', Nonce NA\<rbrace> \<in> parts (spies evs);
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
    69
        Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs)\<rbrakk>  
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
    70
       \<Longrightarrow> Nonce NA \<in> analz (spies evs)"
76291
616405057951 Trying to clean up some messy proofs
paulson <lp15@cam.ac.uk>
parents: 61830
diff changeset
    71
  by (induct rule: ns_public.induct) (auto intro: analz_insertI)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    72
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    73
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    74
text \<open>Unicity for {term NS1}: nonce {term NA} identifies agents {term A} and {term B}\<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    75
lemma unique_NA: 
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    76
  assumes NA: "Crypt(pubEK B)  \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    77
              "Crypt(pubEK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    78
              "Nonce NA \<notin> analz (spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    79
    and evs: "evs \<in> ns_public"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    80
  shows "A=A' \<and> B=B'"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    81
  using evs NA
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    82
  by (induction rule: ns_public.induct) (auto intro!: analz_insertI split: if_split_asm)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    83
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    84
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    85
text \<open>Secrecy: Spy does not see the nonce sent in msg {term NS1} if {term A} and {term B} are secure
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    86
  The major premise "Says A B ..." makes it a dest-rule, hence the given assumption order. \<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    87
theorem Spy_not_see_NA: 
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    88
  assumes NA: "Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    89
              "A \<notin> bad" "B \<notin> bad"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    90
    and evs: "evs \<in> ns_public"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    91
  shows "Nonce NA \<notin> analz (spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    92
  using evs NA
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    93
proof (induction rule: ns_public.induct)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    94
  case (Fake evsf X B)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    95
  then show ?case
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    96
    by spy_analz
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    97
next
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    98
  case (NS2 evs2 NB A' B NA A)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
    99
  then show ?case
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   100
    by simp (metis Says_imp_analz_Spy analz_into_parts parts.simps unique_NA usedI)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   101
next
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   102
  case (NS3 evs3 A B NA B' NB)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   103
  then show ?case
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   104
    by simp (meson Says_imp_analz_Spy analz_into_parts no_nonce_NS1_NS2)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   105
qed auto
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   106
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   107
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   108
text \<open>Authentication for {term A}: if she receives message 2 and has used {term NA}
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   109
  to start a run, then {term B} has sent message 2.\<close>
76368
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   110
lemma A_trusts_NS2_lemma: 
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   111
    "\<lbrakk>evs \<in> ns_public;            
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   112
      Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs);
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   113
      Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   114
      A \<notin> bad; B \<notin> bad\<rbrakk>
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   115
     \<Longrightarrow> Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   116
  by (induct rule: ns_public.induct) (auto dest: Spy_not_see_NA unique_NA)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   117
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   118
theorem A_trusts_NS2: 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   119
     "\<lbrakk>Says A  B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;   
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   120
       Says B' A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs;
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   121
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   122
      \<Longrightarrow> Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   123
  by (blast intro: A_trusts_NS2_lemma)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   124
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   125
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   126
text \<open>If the encrypted message appears then it originated with Alice in {term NS1}\<close>
76368
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   127
lemma B_trusts_NS1:
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   128
     "\<lbrakk>evs \<in> ns_public;                                     
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   129
       Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs);
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   130
       Nonce NA \<notin> analz (spies evs)\<rbrakk>
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   131
      \<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   132
  by (induct evs rule: ns_public.induct) (use analz_insertI in \<open>auto split: if_split_asm\<close>)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   133
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   134
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   135
subsection \<open>Authenticity properties obtained from {term NS2}\<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   136
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   137
text \<open>Unicity for {term NS2}: nonce {term NB} identifies nonce {term NA} and agent {term A} 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   138
  [proof closely follows that for @{thm [source] unique_NA}]\<close>
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   139
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   140
lemma unique_NB [dest]: 
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   141
  assumes NB: "Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts(spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   142
              "Crypt(pubEK A') \<lbrace>Nonce NA', Nonce NB\<rbrace> \<in> parts(spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   143
              "Nonce NB \<notin> analz (spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   144
    and evs: "evs \<in> ns_public"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   145
  shows "A=A' \<and> NA=NA'"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   146
  using evs NB 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   147
  by (induction rule: ns_public.induct) (auto intro!: analz_insertI split: if_split_asm)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   148
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   149
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   150
text \<open>{term NB} remains secret \emph{provided} Alice never responds with round 3\<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   151
theorem Spy_not_see_NB [dest]:
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   152
  assumes NB: "Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   153
              "\<forall>C. Says A C (Crypt (pubEK C) (Nonce NB)) \<notin> set evs"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   154
              "A \<notin> bad" "B \<notin> bad"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   155
    and evs: "evs \<in> ns_public"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   156
  shows "Nonce NB \<notin> analz (spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   157
  using evs NB evs
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   158
proof (induction rule: ns_public.induct)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   159
  case Fake
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   160
  then show ?case by spy_analz
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   161
next
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   162
  case NS2
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   163
  then show ?case
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   164
    by (auto intro!: no_nonce_NS1_NS2)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   165
qed auto
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   166
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   167
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   168
text \<open>Authentication for {term B}: if he receives message 3 and has used {term NB}
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   169
  in message 2, then {term A} has sent message 3 (to somebody) \<close>
76368
943f99825f39 Replaced some ugly legacy proofs
paulson <lp15@cam.ac.uk>
parents: 76299
diff changeset
   170
lemma B_trusts_NS3_lemma:
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   171
     "\<lbrakk>evs \<in> ns_public; 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   172
       Crypt (pubEK B) (Nonce NB) \<in> parts (spies evs); 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   173
       Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs; 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   174
       A \<notin> bad;  B \<notin> bad\<rbrakk>                    
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   175
      \<Longrightarrow> \<exists>C. Says A C (Crypt (pubEK C) (Nonce NB)) \<in> set evs"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   176
proof (induction rule: ns_public.induct)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   177
  case (NS3 evs3 A B NA B' NB)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   178
  then show ?case
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   179
    by simp (blast intro: no_nonce_NS1_NS2)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   180
qed auto
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   181
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   182
theorem B_trusts_NS3:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   183
     "\<lbrakk>Says B A  (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs;
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   184
       Says A' B (Crypt (pubEK B) (Nonce NB)) \<in> set evs;             
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   185
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                    
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   186
      \<Longrightarrow> \<exists>C. Says A C (Crypt (pubEK C) (Nonce NB)) \<in> set evs"
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   187
  by (blast intro: B_trusts_NS3_lemma)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   188
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   189
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   190
text \<open>Can we strengthen the secrecy theorem @{thm[source]Spy_not_see_NB}?  NO\<close>
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   191
lemma "\<lbrakk>evs \<in> ns_public; 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   192
        Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs; 
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   193
        A \<notin> bad; B \<notin> bad\<rbrakk>            
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   194
       \<Longrightarrow> Nonce NB \<notin> analz (spies evs)"
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   195
apply (induction rule: ns_public.induct, simp_all, spy_analz)
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   196
(*{term NS1}: by freshness*)
11150
67387142225e Streamlining for the bug fix in Blast.
paulson
parents: 11104
diff changeset
   197
apply blast
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   198
(*{term NS2}: by freshness and unicity of {term NB}*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   199
apply (blast intro: no_nonce_NS1_NS2)
76297
e7f9e5b3a36a tidying of some old proofs
paulson <lp15@cam.ac.uk>
parents: 76291
diff changeset
   200
(*{term NS3}: unicity of {term NB} identifies {term A} and {term NA}, but not {term B}*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   201
apply clarify
13507
febb8e5d2a9d tidying of Isar scripts
paulson
parents: 11366
diff changeset
   202
apply (frule_tac A' = A in 
febb8e5d2a9d tidying of Isar scripts
paulson
parents: 11366
diff changeset
   203
       Says_imp_knows_Spy [THEN parts.Inj, THEN unique_NB], auto)
38964
b1a7bef0907a renewing specifications in HOL-Auth
bulwahn
parents: 37936
diff changeset
   204
apply (rename_tac evs3 B' C)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   205
txt\<open>This is the attack!
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
   206
@{subgoals[display,indent=0,margin=65]}
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   207
\<close>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   208
oops
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   209
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   210
(*
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   211
THIS IS THE ATTACK!
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   212
Level 8
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   213
!!evs. \<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   214
       \<Longrightarrow> Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs \<longrightarrow>
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   215
           Nonce NB \<notin> analz (spies evs)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   216
 1. !!C B' evs3.
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   217
       \<lbrakk>A \<notin> bad; B \<notin> bad; evs3 \<in> ns_public
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   218
        Says A C (Crypt (pubEK C) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   219
        Says B' A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs3; 
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   220
        C \<in> bad;
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   221
        Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs3;
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   222
        Nonce NB \<notin> analz (spies evs3)\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   223
       \<Longrightarrow> False
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   224
*)
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   225
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   226
end