src/HOL/Auth/Kerberos_BAN.thy
author wenzelm
Sat, 04 Nov 2017 12:39:25 +0100
changeset 67001 b34fbf33a7ea
parent 61830 4f5ab843cf5b
child 67613 ce654b0e6d69
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 36866
diff changeset
     1
(*  Title:      HOL/Auth/Kerberos_BAN.thy
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     2
    Author:     Giampaolo Bella, Cambridge University Computer Laboratory
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     3
    Copyright   1998  University of Cambridge
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     4
*)
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     5
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
     6
section\<open>The Kerberos Protocol, BAN Version\<close>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14207
diff changeset
     8
theory Kerberos_BAN imports Public begin
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     9
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    10
text\<open>From page 251 of
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    11
  Burrows, Abadi and Needham (1989).  A Logic of Authentication.
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    12
  Proc. Royal Soc. 426
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    13
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    14
  Confidentiality (secrecy) and authentication properties are also
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    15
  given in a termporal version: strong guarantees in a little abstracted 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    16
  - but very realistic - model.
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    17
\<close>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    18
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    19
(* Temporal model of accidents: session keys can be leaked
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    20
                                ONLY when they have expired *)
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    21
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    22
consts
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    23
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    24
    (*Duration of the session key*)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    25
    sesKlife   :: nat
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    26
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    27
    (*Duration of the authenticator*)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    28
    authlife :: nat
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    29
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    30
text\<open>The ticket should remain fresh for two journeys on the network at least\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    31
specification (sesKlife)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    32
  sesKlife_LB [iff]: "2 \<le> sesKlife"
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13926
diff changeset
    33
    by blast
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    34
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    35
text\<open>The authenticator only for one journey\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    36
specification (authlife)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
    37
  authlife_LB [iff]:    "authlife \<noteq> 0"
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13926
diff changeset
    38
    by blast
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13926
diff changeset
    39
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 18886
diff changeset
    40
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    41
  CT :: "event list=>nat" where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 18886
diff changeset
    42
  "CT == length "
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    43
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    44
abbreviation
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    45
  expiredK :: "[nat, event list] => bool" where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 18886
diff changeset
    46
  "expiredK T evs == sesKlife + T < CT evs"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    47
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    48
abbreviation
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    49
  expiredA :: "[nat, event list] => bool" where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 18886
diff changeset
    50
  "expiredA T evs == authlife + T < CT evs"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    51
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    52
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    53
definition
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    54
 (* A is the true creator of X if she has sent X and X never appeared on
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    55
    the trace before this event. Recall that traces grow from head. *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    56
  Issues :: "[agent, agent, msg, event list] => bool"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    57
             ("_ Issues _ with _ on _") where
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    58
   "A Issues B with X on evs =
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    59
      (\<exists>Y. Says A B Y \<in> set evs & X \<in> parts {Y} &
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    60
        X \<notin> parts (spies (takeWhile (% z. z  \<noteq> Says A B Y) (rev evs))))"
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    61
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    62
definition
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    63
 (* Yields the subtrace of a given trace from its beginning to a given event *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    64
  before :: "[event, event list] => event list" ("before _ on _")
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    65
  where "before ev on evs = takeWhile (% z. z ~= ev) (rev evs)"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    66
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    67
definition
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    68
 (* States than an event really appears only once on a trace *)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    69
  Unique :: "[event, event list] => bool" ("Unique _ on _")
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32960
diff changeset
    70
  where "Unique ev on evs = (ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs)))"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    71
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    72
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    73
inductive_set bankerberos :: "event list set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    74
 where
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    75
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    76
   Nil:  "[] \<in> bankerberos"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    77
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    78
 | Fake: "\<lbrakk> evsf \<in> bankerberos;  X \<in> synth (analz (spies evsf)) \<rbrakk>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    79
          \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    80
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    81
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    82
 | BK1:  "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    83
          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    84
                \<in>  bankerberos"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    85
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
    86
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    87
 | BK2:  "\<lbrakk> evs2 \<in> bankerberos;  Key K \<notin> used evs2; K \<in> symKeys;
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    88
             Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    89
          \<Longrightarrow> Says Server A
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    90
                (Crypt (shrK A)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    91
                   \<lbrace>Number (CT evs2), Agent B, Key K,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    92
                    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    93
                # evs2 \<in> bankerberos"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    94
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    95
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    96
 | BK3:  "\<lbrakk> evs3 \<in> bankerberos;
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    97
             Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    98
               \<in> set evs3;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
    99
             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   100
             \<not> expiredK Tk evs3 \<rbrakk>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   101
          \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   102
               # evs3 \<in> bankerberos"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   103
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   104
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   105
 | BK4:  "\<lbrakk> evs4 \<in> bankerberos;
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   106
             Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   107
                         (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   108
             \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   109
          \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   110
                \<in> bankerberos"
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   111
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   112
        (*Old session keys may become compromised*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   113
 | Oops: "\<lbrakk> evso \<in> bankerberos;
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   114
         Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   115
               \<in> set evso;
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   116
             expiredK Tk evso \<rbrakk>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32406
diff changeset
   117
          \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   118
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   119
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   120
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14126
diff changeset
   121
declare parts.Body [dest]
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14126
diff changeset
   122
declare analz_into_parts [dest]
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14126
diff changeset
   123
declare Fake_parts_insert_in_Un [dest]
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   124
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   125
text\<open>A "possibility property": there are traces that reach the end.\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   126
lemma "\<lbrakk>Key K \<notin> used []; K \<in> symKeys\<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   127
       \<Longrightarrow> \<exists>Timestamp. \<exists>evs \<in> bankerberos.
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   128
             Says B A (Crypt K (Number Timestamp))
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   129
                  \<in> set evs"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   130
apply (cut_tac sesKlife_LB)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   131
apply (intro exI bexI)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   132
apply (rule_tac [2]
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   133
           bankerberos.Nil [THEN bankerberos.BK1, THEN bankerberos.BK2,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   134
                             THEN bankerberos.BK3, THEN bankerberos.BK4])
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   135
apply (possibility, simp_all (no_asm_simp) add: used_Cons)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   136
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   137
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   138
subsection\<open>Lemmas for reasoning about predicate "Issues"\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   139
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   140
lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   141
apply (induct_tac "evs")
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   142
apply (rename_tac [2] a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   143
apply (induct_tac [2] "a", auto)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   144
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   145
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   146
lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   147
apply (induct_tac "evs")
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   148
apply (rename_tac [2] a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   149
apply (induct_tac [2] "a", auto)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   150
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   151
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   152
lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   153
          (if A:bad then insert X (spies evs) else spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   154
apply (induct_tac "evs")
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   155
apply (rename_tac [2] a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   156
apply (induct_tac [2] "a", auto)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   157
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   158
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   159
lemma spies_evs_rev: "spies evs = spies (rev evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   160
apply (induct_tac "evs")
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   161
apply (rename_tac [2] a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   162
apply (induct_tac [2] "a")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   163
apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   164
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   165
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   166
lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono]
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   167
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   168
lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   169
apply (induct_tac "evs")
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   170
apply (rename_tac [2] a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   171
apply (induct_tac [2] "a", auto)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   172
txt\<open>Resembles \<open>used_subset_append\<close> in theory Event.\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   173
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   174
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   175
lemmas parts_spies_takeWhile_mono = spies_takeWhile [THEN parts_mono]
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   176
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   177
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   178
text\<open>Lemmas for reasoning about predicate "before"\<close>
39251
8756b44582e2 Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
paulson
parents: 37936
diff changeset
   179
lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   180
apply (induct_tac "evs")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   181
apply simp
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   182
apply (rename_tac a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   183
apply (induct_tac "a")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   184
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   185
done
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   186
39251
8756b44582e2 Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
paulson
parents: 37936
diff changeset
   187
lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   188
apply (induct_tac "evs")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   189
apply simp
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   190
apply (rename_tac a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   191
apply (induct_tac "a")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   192
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   193
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   194
39251
8756b44582e2 Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
paulson
parents: 37936
diff changeset
   195
lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   196
apply (induct_tac "evs")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   197
apply simp
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   198
apply (rename_tac a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   199
apply (induct_tac "a")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   200
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   201
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   202
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   203
lemma used_evs_rev: "used evs = used (rev evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   204
apply (induct_tac "evs")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   205
apply simp
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   206
apply (rename_tac a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   207
apply (induct_tac "a")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   208
apply (simp add: used_Says_rev)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   209
apply (simp add: used_Gets_rev)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   210
apply (simp add: used_Notes_rev)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   211
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   212
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   213
lemma used_takeWhile_used [rule_format]: 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   214
      "x : used (takeWhile P X) --> x : used X"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   215
apply (induct_tac "X")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   216
apply simp
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 42749
diff changeset
   217
apply (rename_tac a b)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   218
apply (induct_tac "a")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   219
apply (simp_all add: used_Nil)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   220
apply (blast dest!: initState_into_used)+
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   221
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   222
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   223
lemma set_evs_rev: "set evs = set (rev evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   224
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   225
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   226
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   227
lemma takeWhile_void [rule_format]:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   228
      "x \<notin> set evs \<longrightarrow> takeWhile (\<lambda>z. z \<noteq> x) evs = evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   229
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   230
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   231
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   232
(**** Inductive proofs about bankerberos ****)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   233
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   234
text\<open>Forwarding Lemma for reasoning about the encrypted portion of message BK3\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   235
lemma BK3_msg_in_parts_spies:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   236
     "Says S A (Crypt KA \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   237
      \<Longrightarrow> X \<in> parts (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   238
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   239
done
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   240
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   241
lemma Oops_parts_spies:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   242
     "Says Server A (Crypt (shrK A) \<lbrace>Timestamp, B, K, X\<rbrace>) \<in> set evs
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   243
      \<Longrightarrow> K \<in> parts (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   244
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   245
done
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   246
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   247
text\<open>Spy never sees another agent's shared key! (unless it's bad at start)\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   248
lemma Spy_see_shrK [simp]:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   249
     "evs \<in> bankerberos \<Longrightarrow> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   250
apply (erule bankerberos.induct)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   251
apply (frule_tac [7] Oops_parts_spies)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   252
apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast+)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   253
done
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   254
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   255
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   256
lemma Spy_analz_shrK [simp]:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   257
     "evs \<in> bankerberos \<Longrightarrow> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   258
apply auto
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   259
done
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   260
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   261
lemma Spy_see_shrK_D [dest!]:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   262
     "\<lbrakk> Key (shrK A) \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   263
                evs \<in> bankerberos \<rbrakk> \<Longrightarrow> A:bad"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   264
apply (blast dest: Spy_see_shrK)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   265
done
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   266
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   267
lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D,  dest!]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   268
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   269
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   270
text\<open>Nobody can have used non-existent keys!\<close>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   271
lemma new_keys_not_used [simp]:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   272
    "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> bankerberos\<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   273
     \<Longrightarrow> K \<notin> keysFor (parts (spies evs))"
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   274
apply (erule rev_mp)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   275
apply (erule bankerberos.induct)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   276
apply (frule_tac [7] Oops_parts_spies)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   277
apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   278
txt\<open>Fake\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   279
apply (force dest!: keysFor_parts_insert)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   280
txt\<open>BK2, BK3, BK4\<close>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   281
apply (force dest!: analz_shrK_Decrypt)+
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   282
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   283
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   284
subsection\<open>Lemmas concerning the form of items passed in messages\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   285
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   286
text\<open>Describes the form of K, X and K' when the Server sends this message.\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   287
lemma Says_Server_message_form:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   288
     "\<lbrakk> Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   289
         \<in> set evs; evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   290
      \<Longrightarrow> K' = shrK A & K \<notin> range shrK &
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   291
          Ticket = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>) &
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   292
          Key K \<notin> used(before
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   293
                  Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   294
                  on evs) &
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   295
          Tk = CT(before 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   296
                  Says Server A (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   297
                  on evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   298
apply (unfold before_def)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   299
apply (erule rev_mp)
32406
9ea59bd1397a Simplified some proofs using metis.
paulson
parents: 25134
diff changeset
   300
apply (erule bankerberos.induct, simp_all add: takeWhile_tail)
42749
47f283fcf2ae Metis doesn't find an old proof in acceptable time now that higher-order equality reasoning is supported -- tuned proof script to help it
blanchet
parents: 39251
diff changeset
   301
apply auto
47f283fcf2ae Metis doesn't find an old proof in acceptable time now that higher-order equality reasoning is supported -- tuned proof script to help it
blanchet
parents: 39251
diff changeset
   302
 apply (metis length_rev set_rev takeWhile_void used_evs_rev)+
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   303
done
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   304
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   305
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   306
text\<open>If the encrypted message appears then it originated with the Server
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   307
  PROVIDED that A is NOT compromised!
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   308
  This allows A to verify freshness of the session key.
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   309
\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   310
lemma Kab_authentic:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   311
     "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   312
           \<in> parts (spies evs);
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   313
         A \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   314
       \<Longrightarrow> Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   315
             \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   316
apply (erule rev_mp)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   317
apply (erule bankerberos.induct)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   318
apply (frule_tac [7] Oops_parts_spies)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   319
apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   320
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   321
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   322
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   323
text\<open>If the TICKET appears then it originated with the Server\<close>
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   324
text\<open>FRESHNESS OF THE SESSION KEY to B\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   325
lemma ticket_authentic:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   326
     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   327
         B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   328
       \<Longrightarrow> Says Server A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   329
            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   330
                          Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   331
           \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   332
apply (erule rev_mp)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   333
apply (erule bankerberos.induct)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   334
apply (frule_tac [7] Oops_parts_spies)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   335
apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all, blast)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   336
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   337
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   338
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   339
text\<open>EITHER describes the form of X when the following message is sent,
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   340
  OR     reduces it to the Fake case.
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   341
  Use \<open>Says_Server_message_form\<close> if applicable.\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   342
lemma Says_S_message_form:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   343
     "\<lbrakk> Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   344
            \<in> set evs;
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   345
         evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   346
 \<Longrightarrow> (K \<notin> range shrK & X = (Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>))
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   347
          | X \<in> analz (spies evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   348
apply (case_tac "A \<in> bad")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   349
apply (force dest!: Says_imp_spies [THEN analz.Inj])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   350
apply (frule Says_imp_spies [THEN parts.Inj])
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   351
apply (blast dest!: Kab_authentic Says_Server_message_form)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   352
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   353
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   354
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   355
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   356
(****
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   357
 The following is to prove theorems of the form
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   358
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   359
  Key K \<in> analz (insert (Key KAB) (spies evs)) \<Longrightarrow>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   360
  Key K \<in> analz (spies evs)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   361
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   362
 A more general formula must be proved inductively.
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   363
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   364
****)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   365
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   366
text\<open>Session keys are not used to encrypt other session keys\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   367
lemma analz_image_freshK [rule_format (no_asm)]:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   368
     "evs \<in> bankerberos \<Longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   369
   \<forall>K KK. KK \<subseteq> - (range shrK) \<longrightarrow>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   370
          (Key K \<in> analz (Key`KK Un (spies evs))) =
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   371
          (K \<in> KK | Key K \<in> analz (spies evs))"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   372
apply (erule bankerberos.induct)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   373
apply (drule_tac [7] Says_Server_message_form)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   374
apply (erule_tac [5] Says_S_message_form [THEN disjE], analz_freshK, spy_analz, auto) 
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   375
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   376
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   377
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   378
lemma analz_insert_freshK:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   379
     "\<lbrakk> evs \<in> bankerberos;  KAB \<notin> range shrK \<rbrakk> \<Longrightarrow>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   380
      (Key K \<in> analz (insert (Key KAB) (spies evs))) =
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   381
      (K = KAB | Key K \<in> analz (spies evs))"
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   382
apply (simp only: analz_image_freshK analz_image_freshK_simps)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   383
done
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   384
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   385
text\<open>The session key K uniquely identifies the message\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   386
lemma unique_session_keys:
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   387
     "\<lbrakk> Says Server A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   388
           (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>) \<in> set evs;
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   389
         Says Server A'
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   390
          (Crypt (shrK A') \<lbrace>Number Tk', Agent B', Key K, X'\<rbrace>) \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   391
         evs \<in> bankerberos \<rbrakk> \<Longrightarrow> A=A' & Tk=Tk' & B=B' & X = X'"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   392
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   393
apply (erule rev_mp)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   394
apply (erule bankerberos.induct)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   395
apply (frule_tac [7] Oops_parts_spies)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   396
apply (frule_tac [5] BK3_msg_in_parts_spies, simp_all)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   397
txt\<open>BK2: it can't be a new key\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   398
apply blast
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   399
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   400
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   401
lemma Server_Unique:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   402
     "\<lbrakk> Says Server A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   403
            (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   404
        evs \<in> bankerberos \<rbrakk> \<Longrightarrow> 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   405
   Unique Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   406
   on evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   407
apply (erule rev_mp, erule bankerberos.induct, simp_all add: Unique_def)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   408
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   409
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   410
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   411
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   412
subsection\<open>Non-temporal guarantees, explicitly relying on non-occurrence of
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   413
oops events - refined below by temporal guarantees\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   414
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   415
text\<open>Non temporal treatment of confidentiality\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   416
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   417
text\<open>Lemma: the session key sent in msg BK2 would be lost by oops
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   418
    if the spy could see it!\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   419
lemma lemma_conf [rule_format (no_asm)]:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   420
     "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   421
  \<Longrightarrow> Says Server A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   422
          (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   423
                            Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   424
         \<in> set evs \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   425
      Key K \<in> analz (spies evs) \<longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   426
apply (erule bankerberos.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   427
apply (frule_tac [7] Says_Server_message_form)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   428
apply (frule_tac [5] Says_S_message_form [THEN disjE])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   429
apply (simp_all (no_asm_simp) add: analz_insert_eq analz_insert_freshK pushes)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   430
txt\<open>Fake\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   431
apply spy_analz
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   432
txt\<open>BK2\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   433
apply (blast intro: parts_insertI)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   434
txt\<open>BK3\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   435
apply (case_tac "Aa \<in> bad")
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   436
 prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   437
apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   438
txt\<open>Oops\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   439
apply (blast dest: unique_session_keys)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   440
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   441
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   442
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   443
text\<open>Confidentiality for the Server: Spy does not see the keys sent in msg BK2
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   444
as long as they have not expired.\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   445
lemma Confidentiality_S:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   446
     "\<lbrakk> Says Server A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   447
          (Crypt K' \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>) \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   448
        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   449
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   450
      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   451
apply (frule Says_Server_message_form, assumption)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   452
apply (blast intro: lemma_conf)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   453
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   454
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   455
text\<open>Confidentiality for Alice\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   456
lemma Confidentiality_A:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   457
     "\<lbrakk> Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   458
        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   459
        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   460
      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   461
apply (blast dest!: Kab_authentic Confidentiality_S)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   462
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   463
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   464
text\<open>Confidentiality for Bob\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   465
lemma Confidentiality_B:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   466
     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   467
          \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   468
        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   469
        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   470
      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   471
apply (blast dest!: ticket_authentic Confidentiality_S)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   472
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   473
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   474
text\<open>Non temporal treatment of authentication\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   475
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   476
text\<open>Lemmas \<open>lemma_A\<close> and \<open>lemma_B\<close> in fact are common to both temporal and non-temporal treatments\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   477
lemma lemma_A [rule_format]:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   478
     "\<lbrakk> A \<notin> bad; B \<notin> bad; evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   479
      \<Longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   480
         Key K \<notin> analz (spies evs) \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   481
         Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   482
         \<in> set evs \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   483
          Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs) \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   484
         Says A B \<lbrace>X, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   485
             \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   486
apply (erule bankerberos.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   487
apply (frule_tac [7] Oops_parts_spies)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   488
apply (frule_tac [5] Says_S_message_form)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   489
apply (frule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   490
apply (simp_all (no_asm_simp) add: all_conj_distrib)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   491
txt\<open>Fake\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   492
apply blast
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   493
txt\<open>BK2\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   494
apply (force dest: Crypt_imp_invKey_keysFor)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   495
txt\<open>BK3\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   496
apply (blast dest: Kab_authentic unique_session_keys)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   497
done
32406
9ea59bd1397a Simplified some proofs using metis.
paulson
parents: 25134
diff changeset
   498
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   499
lemma lemma_B [rule_format]:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   500
     "\<lbrakk> B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   501
      \<Longrightarrow> Key K \<notin> analz (spies evs) \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   502
          Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   503
          \<in> set evs \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   504
          Crypt K (Number Ta) \<in> parts (spies evs) \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   505
          Says B A (Crypt K (Number Ta)) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   506
apply (erule bankerberos.induct)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   507
apply (frule_tac [7] Oops_parts_spies)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   508
apply (frule_tac [5] Says_S_message_form)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   509
apply (drule_tac [6] BK3_msg_in_parts_spies, analz_mono_contra)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   510
apply (simp_all (no_asm_simp) add: all_conj_distrib)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   511
txt\<open>Fake\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   512
apply blast
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   513
txt\<open>BK2\<close> 
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   514
apply (force dest: Crypt_imp_invKey_keysFor)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   515
txt\<open>BK4\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   516
apply (blast dest: ticket_authentic unique_session_keys
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   517
                   Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   518
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   519
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   520
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   521
text\<open>The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   522
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   523
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   524
text\<open>Authentication of A to B\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   525
lemma B_authenticates_A_r:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   526
     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   527
         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   528
        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   529
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   530
      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   531
                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   532
apply (blast dest!: ticket_authentic
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   533
          intro!: lemma_A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   534
          elim!: Confidentiality_S [THEN [2] rev_notE])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   535
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   536
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   537
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   538
text\<open>Authentication of B to A\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   539
lemma A_authenticates_B_r:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   540
     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   541
        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   542
        Notes Spy \<lbrace>Number Tk, Key K\<rbrace> \<notin> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   543
        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   544
      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   545
apply (blast dest!: Kab_authentic
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   546
          intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   547
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   548
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   549
lemma B_authenticates_A:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   550
     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   551
         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   552
        Key K \<notin> analz (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   553
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   554
      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   555
                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   556
apply (blast dest!: ticket_authentic intro!: lemma_A)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   557
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   558
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   559
lemma A_authenticates_B:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   560
     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   561
        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   562
        Key K \<notin> analz (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   563
        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   564
      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   565
apply (blast dest!: Kab_authentic intro!: lemma_B)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   566
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   567
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   568
subsection\<open>Temporal guarantees, relying on a temporal check that insures that
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   569
no oops event occurred. These are available in the sense of goal availability\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   570
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   571
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   572
text\<open>Temporal treatment of confidentiality\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   573
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   574
text\<open>Lemma: the session key sent in msg BK2 would be EXPIRED
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   575
    if the spy could see it!\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   576
lemma lemma_conf_temporal [rule_format (no_asm)]:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   577
     "\<lbrakk> A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   578
  \<Longrightarrow> Says Server A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   579
          (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   580
                            Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>\<rbrace>)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   581
         \<in> set evs \<longrightarrow>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   582
      Key K \<in> analz (spies evs) \<longrightarrow> expiredK Tk evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   583
apply (erule bankerberos.induct)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   584
apply (frule_tac [7] Says_Server_message_form)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   585
apply (frule_tac [5] Says_S_message_form [THEN disjE])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   586
apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   587
txt\<open>Fake\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   588
apply spy_analz
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   589
txt\<open>BK2\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   590
apply (blast intro: parts_insertI less_SucI)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   591
txt\<open>BK3\<close>
32406
9ea59bd1397a Simplified some proofs using metis.
paulson
parents: 25134
diff changeset
   592
apply (metis Crypt_Spy_analz_bad Kab_authentic Says_imp_analz_Spy 
9ea59bd1397a Simplified some proofs using metis.
paulson
parents: 25134
diff changeset
   593
          Says_imp_parts_knows_Spy analz.Snd less_SucI unique_session_keys)
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   594
txt\<open>Oops: PROOF FAILS if unsafe intro below\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   595
apply (blast dest: unique_session_keys intro!: less_SucI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   596
done
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   597
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   598
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   599
text\<open>Confidentiality for the Server: Spy does not see the keys sent in msg BK2
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   600
as long as they have not expired.\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   601
lemma Confidentiality_S_temporal:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   602
     "\<lbrakk> Says Server A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   603
          (Crypt K' \<lbrace>Number T, Agent B, Key K, X\<rbrace>) \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   604
         \<not> expiredK T evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   605
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   606
      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   607
apply (frule Says_Server_message_form, assumption)
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   608
apply (blast intro: lemma_conf_temporal)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   609
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   610
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   611
text\<open>Confidentiality for Alice\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   612
lemma Confidentiality_A_temporal:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   613
     "\<lbrakk> Crypt (shrK A) \<lbrace>Number T, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   614
         \<not> expiredK T evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   615
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   616
      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   617
apply (blast dest!: Kab_authentic Confidentiality_S_temporal)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   618
done
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   619
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   620
text\<open>Confidentiality for Bob\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   621
lemma Confidentiality_B_temporal:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   622
     "\<lbrakk> Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   623
          \<in> parts (spies evs);
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   624
        \<not> expiredK Tk evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   625
        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   626
      \<rbrakk> \<Longrightarrow> Key K \<notin> analz (spies evs)"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   627
apply (blast dest!: ticket_authentic Confidentiality_S_temporal)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   628
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   629
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   630
text\<open>Temporal treatment of authentication\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   631
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   632
text\<open>Authentication of A to B\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   633
lemma B_authenticates_A_temporal:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   634
     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   635
         Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   636
         \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   637
         \<not> expiredK Tk evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   638
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   639
      \<Longrightarrow> Says A B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   640
                     Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   641
apply (blast dest!: ticket_authentic
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   642
          intro!: lemma_A
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   643
          elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   644
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   645
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   646
text\<open>Authentication of B to A\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   647
lemma A_authenticates_B_temporal:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   648
     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   649
         Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   650
         \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   651
         \<not> expiredK Tk evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   652
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   653
      \<Longrightarrow> Says B A (Crypt K (Number Ta)) \<in> set evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   654
apply (blast dest!: Kab_authentic
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   655
          intro!: lemma_B elim!: Confidentiality_S_temporal [THEN [2] rev_notE])
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   656
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   657
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   658
subsection\<open>Treatment of the key distribution goal using trace inspection. All
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   659
guarantees are in non-temporal form, hence non available, though their temporal
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   660
form is trivial to derive. These guarantees also convey a stronger form of 
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   661
authentication - non-injective agreement on the session key\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   662
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   663
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   664
lemma B_Issues_A:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   665
     "\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   666
         Key K \<notin> analz (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   667
         A \<notin> bad;  B \<notin> bad; evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   668
      \<Longrightarrow> B Issues A with (Crypt K (Number Ta)) on evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   669
apply (simp (no_asm) add: Issues_def)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   670
apply (rule exI)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   671
apply (rule conjI, assumption)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   672
apply (simp (no_asm))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   673
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   674
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   675
apply (erule bankerberos.induct, analz_mono_contra)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   676
apply (simp_all (no_asm_simp))
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   677
txt\<open>fake\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   678
apply blast
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   679
txt\<open>K4 obviously is the non-trivial case\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   680
apply (simp add: takeWhile_tail)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   681
apply (blast dest: ticket_authentic parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] intro: A_authenticates_B_temporal)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   682
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   683
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   684
lemma A_authenticates_and_keydist_to_B:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   685
     "\<lbrakk> Crypt K (Number Ta) \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   686
        Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, X\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   687
         Key K \<notin> analz (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   688
         A \<notin> bad;  B \<notin> bad; evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   689
      \<Longrightarrow> B Issues A with (Crypt K (Number Ta)) on evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   690
apply (blast dest!: A_authenticates_B B_Issues_A)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   691
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   692
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   693
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   694
lemma A_Issues_B:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   695
     "\<lbrakk> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   696
           \<in> set evs;
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   697
         Key K \<notin> analz (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   698
         A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   699
   \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) on evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   700
apply (simp (no_asm) add: Issues_def)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   701
apply (rule exI)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   702
apply (rule conjI, assumption)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   703
apply (simp (no_asm))
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   704
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   705
apply (erule rev_mp)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   706
apply (erule bankerberos.induct, analz_mono_contra)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   707
apply (simp_all (no_asm_simp))
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   708
txt\<open>fake\<close>
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   709
apply blast
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   710
txt\<open>K3 is the non trivial case\<close>
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   711
apply (simp add: takeWhile_tail)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   712
apply auto (*Technically unnecessary, merely clarifies the subgoal as it is presemted in the book*)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   713
apply (blast dest: Kab_authentic Says_Server_message_form parts_spies_takeWhile_mono [THEN subsetD] parts_spies_evs_revD2 [THEN subsetD] 
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   714
             intro!: B_authenticates_A)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   715
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   716
18886
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   717
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   718
lemma B_authenticates_and_keydist_to_A:
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   719
     "\<lbrakk> Crypt K \<lbrace>Agent A, Number Ta\<rbrace> \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   720
        Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>  \<in> parts (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   721
        Key K \<notin> analz (spies evs);
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   722
        A \<notin> bad;  B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   723
   \<Longrightarrow> A Issues B with (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) on evs"
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   724
apply (blast dest: B_authenticates_A A_Issues_B)
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   725
done
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   726
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   727
9f27383426db new and updated protocol proofs by Giamp Bella
paulson
parents: 16417
diff changeset
   728
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   729
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   730
end