src/HOL/Auth/Kerberos_BAN.thy
author berghofe
Fri, 11 Jul 2003 14:55:17 +0200
changeset 14102 8af7334af4b3
parent 13926 6e62e5357a10
child 14126 28824746d046
permissions -rw-r--r--
- Installed specific code generator for equality enforcing that arguments do not have function types, which would result in an error message during compilation. - Added test case generators for basic types.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Kerberos_BAN
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     2
    ID:         $Id$
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     3
    Author:     Giampaolo Bella, Cambridge University Computer Laboratory
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     5
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     6
The Kerberos protocol, BAN version.
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     7
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     8
From page 251 of
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
     9
  Burrows, Abadi and Needham.  A Logic of Authentication.
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    10
  Proc. Royal Soc. 426 (1989)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    11
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    12
  Confidentiality (secrecy) and authentication properties rely on 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    13
  temporal checks: strong guarantees in a little abstracted - but
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    14
  very realistic - model (see .thy).
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    15
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    16
Tidied and converted to Isar by lcp.
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    17
*)
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    18
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    19
theory Kerberos_BAN = Shared:
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    20
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    21
(* Temporal modelization: session keys can be leaked 
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    22
                          ONLY when they have expired *)
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
syntax
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    25
    CT :: "event list=>nat"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    26
    Expired :: "[nat, event list] => bool"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    27
    RecentAuth :: "[nat, event list] => bool"
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    28
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    29
consts
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    30
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    31
    (*Duration of the session key*)
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    32
    SesKeyLife   :: nat
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    33
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    34
    (*Duration of the authenticator*)
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    35
    AutLife :: nat
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    36
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    37
axioms
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    38
    (*The ticket should remain fresh for two journeys on the network at least*)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    39
    SesKeyLife_LB: "2 <= SesKeyLife"
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    40
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    41
    (*The authenticator only for one journey*)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    42
    AutLife_LB:    "Suc 0 <= AutLife"
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    43
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    44
translations
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    45
   "CT" == "length"
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    46
  
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    47
   "Expired T evs" == "SesKeyLife + T < CT evs"
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    48
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    49
   "RecentAuth T evs" == "CT evs <= AutLife + T"
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    50
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    51
consts  kerberos_ban   :: "event list set"
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    52
inductive "kerberos_ban"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    53
 intros 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    54
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    55
   Nil:  "[] \<in> kerberos_ban"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    56
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    57
   Fake: "[| evsf \<in> kerberos_ban;  X \<in> synth (analz (spies evsf)) |]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    58
	  ==> Says Spy B X # evsf \<in> kerberos_ban"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    59
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    60
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    61
   Kb1:  "[| evs1 \<in> kerberos_ban |]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    62
	  ==> Says A Server {|Agent A, Agent B|} # evs1
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    63
		\<in>  kerberos_ban"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    64
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    65
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    66
   Kb2:  "[| evs2 \<in> kerberos_ban;  Key KAB \<notin> used evs2;
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    67
	     Says A' Server {|Agent A, Agent B|} \<in> set evs2 |]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    68
	  ==> Says Server A 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    69
		(Crypt (shrK A)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    70
		   {|Number (CT evs2), Agent B, Key KAB,  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    71
		    (Crypt (shrK B) {|Number (CT evs2), Agent A, Key KAB|})|}) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    72
		# evs2 \<in> kerberos_ban"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    73
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    74
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    75
   Kb3:  "[| evs3 \<in> kerberos_ban;  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    76
	     Says S A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    77
	       \<in> set evs3;
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    78
	     Says A Server {|Agent A, Agent B|} \<in> set evs3;
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    79
	     ~ Expired Ts evs3 |]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    80
	  ==> Says A B {|X, Crypt K {|Agent A, Number (CT evs3)|} |} 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    81
	       # evs3 \<in> kerberos_ban"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    82
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    83
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    84
   Kb4:  "[| evs4 \<in> kerberos_ban;  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    85
	     Says A' B {|(Crypt (shrK B) {|Number Ts, Agent A, Key K|}), 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    86
			 (Crypt K {|Agent A, Number Ta|}) |}: set evs4;
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    87
	     ~ Expired Ts evs4;  RecentAuth Ta evs4 |]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    88
	  ==> Says B A (Crypt K (Number Ta)) # evs4
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    89
		\<in> kerberos_ban"
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
    90
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    91
	(*Old session keys may become compromised*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    92
   Oops: "[| evso \<in> kerberos_ban;  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    93
	     Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    94
	       \<in> set evso;
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    95
	     Expired Ts evso |]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    96
	  ==> Notes Spy {|Number Ts, Key K|} # evso \<in> kerberos_ban"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    97
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    98
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
    99
declare Says_imp_knows_Spy [THEN parts.Inj, dest] parts.Body [dest]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   100
declare analz_subset_parts [THEN subsetD, dest]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   101
declare Fake_parts_insert [THEN subsetD, dest]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   102
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   103
declare SesKeyLife_LB [iff] AutLife_LB [iff]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   104
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   105
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   106
(*A "possibility property": there are traces that reach the end.*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   107
lemma "\<exists>Timestamp K. \<exists>evs \<in> kerberos_ban.     
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   108
             Says B A (Crypt K (Number Timestamp))  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   109
                  \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   110
apply (cut_tac SesKeyLife_LB)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   111
apply (intro exI bexI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   112
apply (rule_tac [2] 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   113
           kerberos_ban.Nil [THEN kerberos_ban.Kb1, THEN kerberos_ban.Kb2, 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   114
                             THEN kerberos_ban.Kb3, THEN kerberos_ban.Kb4], possibility)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   115
apply (simp_all (no_asm_simp))
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   116
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   117
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
(**** Inductive proofs about kerberos_ban ****)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   120
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   121
(*Forwarding Lemma for reasoning about the encrypted portion of message Kb3*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   122
lemma Kb3_msg_in_parts_spies:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   123
     "Says S A (Crypt KA {|Timestamp, B, K, X|}) \<in> set evs  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   124
      ==> X \<in> parts (spies evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   125
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   126
                              
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   127
lemma Oops_parts_spies:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   128
     "Says Server A (Crypt (shrK A) {|Timestamp, B, K, X|}) \<in> set evs  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   129
      ==> K \<in> parts (spies evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   130
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   131
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   132
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   133
(*Spy never sees another agent's shared key! (unless it's bad at start)*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   134
lemma Spy_see_shrK [simp]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   135
     "evs \<in> kerberos_ban ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   136
apply (erule kerberos_ban.induct) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   137
apply (frule_tac [7] Oops_parts_spies) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   138
apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all)  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   139
apply blast+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   140
done
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   141
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   142
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   143
lemma Spy_analz_shrK [simp]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   144
     "evs \<in> kerberos_ban ==> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   145
apply auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   146
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   147
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   148
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   149
lemma Spy_see_shrK_D [dest!]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   150
     "[| Key (shrK A) \<in> parts (spies evs);        
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   151
                evs \<in> kerberos_ban |] ==> A:bad"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   152
apply (blast dest: Spy_see_shrK)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   153
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   154
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   155
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
   156
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   157
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   158
(*Nobody can have used non-existent keys!*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   159
lemma new_keys_not_used [rule_format, simp]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   160
     "evs \<in> kerberos_ban ==>       
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   161
       Key K \<notin> used evs --> K \<notin> keysFor (parts (spies evs))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   162
apply (erule kerberos_ban.induct) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   163
apply (frule_tac [7] Oops_parts_spies) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   164
apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all)  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   165
(*Fake*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   166
apply (force dest!: keysFor_parts_insert)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   167
(*Kb2, Kb3, Kb4*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   168
apply blast+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   169
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   170
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   171
(** Lemmas concerning the form of items passed in messages **)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   172
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   173
(*Describes the form of K, X and K' when the Server sends this message.*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   174
lemma Says_Server_message_form:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   175
     "[| Says Server A (Crypt K' {|Number Ts, Agent B, Key K, X|})   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   176
         \<in> set evs; evs \<in> kerberos_ban |]                            
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   177
      ==> K \<notin> range shrK &                                          
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   178
          X = (Crypt (shrK B) {|Number Ts, Agent A, Key K|}) &       
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   179
          K' = shrK A"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   180
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   181
apply (erule kerberos_ban.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   182
done
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   183
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   184
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   185
(*If the encrypted message appears then it originated with the Server
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   186
  PROVIDED that A is NOT compromised!
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   187
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   188
  This shows implicitly the FRESHNESS OF THE SESSION KEY to A
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   189
*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   190
lemma A_trusts_K_by_Kb2:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   191
     "[| Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   192
           \<in> parts (spies evs);                           
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   193
         A \<notin> bad;  evs \<in> kerberos_ban |]                 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   194
       ==> Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   195
             \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   196
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   197
apply (erule kerberos_ban.induct) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   198
apply (frule_tac [7] Oops_parts_spies) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   199
apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all)  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   200
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   201
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   202
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   203
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   204
(*If the TICKET appears then it originated with the Server*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   205
(*FRESHNESS OF THE SESSION KEY to B*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   206
lemma B_trusts_K_by_Kb3:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   207
     "[| Crypt (shrK B) {|Number Ts, Agent A, Key K|} \<in> parts (spies evs);  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   208
         B \<notin> bad;  evs \<in> kerberos_ban |]                         
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   209
       ==> Says Server A                                          
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   210
            (Crypt (shrK A) {|Number Ts, Agent B, Key K,                    
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   211
                          Crypt (shrK B) {|Number Ts, Agent A, Key K|}|})   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   212
           \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   213
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   214
apply (erule kerberos_ban.induct) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   215
apply (frule_tac [7] Oops_parts_spies) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   216
apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all)  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   217
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   218
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   219
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   220
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   221
(*EITHER describes the form of X when the following message is sent, 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   222
  OR     reduces it to the Fake case.
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   223
  Use Says_Server_message_form if applicable.*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   224
lemma Says_S_message_form:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   225
     "[| Says S A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})      
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   226
            \<in> set evs;                                                   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   227
         evs \<in> kerberos_ban |]                                           
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   228
 ==> (K \<notin> range shrK & X = (Crypt (shrK B) {|Number Ts, Agent A, Key K|})) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   229
          | X \<in> analz (spies evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   230
apply (case_tac "A \<in> bad")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   231
apply (force dest!: Says_imp_spies [THEN analz.Inj])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   232
apply (frule Says_imp_spies [THEN parts.Inj])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   233
apply (blast dest!: A_trusts_K_by_Kb2 Says_Server_message_form)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   234
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   235
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   236
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   237
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   238
(****
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   239
 The following is to prove theorems of the form
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   240
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   241
  Key K \<in> analz (insert (Key KAB) (spies evs)) ==>
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   242
  Key K \<in> analz (spies evs)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   243
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   244
 A more general formula must be proved inductively.
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   245
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   246
****)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   247
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   248
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   249
(** Session keys are not used to encrypt other session keys **)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   250
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   251
lemma analz_image_freshK [rule_format (no_asm)]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   252
     "evs \<in> kerberos_ban ==>                           
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   253
   \<forall>K KK. KK <= - (range shrK) -->                  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   254
          (Key K \<in> analz (Key`KK Un (spies evs))) =   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   255
          (K \<in> KK | Key K \<in> analz (spies evs))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   256
apply (erule kerberos_ban.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   257
apply (drule_tac [7] Says_Server_message_form)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   258
apply (erule_tac [5] Says_S_message_form [THEN disjE], analz_freshK, spy_analz)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   259
done
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
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   262
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   263
lemma analz_insert_freshK:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   264
     "[| evs \<in> kerberos_ban;  KAB \<notin> range shrK |] ==>      
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   265
      (Key K \<in> analz (insert (Key KAB) (spies evs))) =        
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   266
      (K = KAB | Key K \<in> analz (spies evs))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   267
by (simp only: analz_image_freshK analz_image_freshK_simps)
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
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   270
(** The session key K uniquely identifies the message **)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   271
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   272
lemma unique_session_keys:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   273
     "[| Says Server A                                     
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   274
           (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}) \<in> set evs;   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   275
         Says Server A'                                    
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   276
          (Crypt (shrK A') {|Number Ts', Agent B', Key K, X'|}) \<in> set evs; 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   277
         evs \<in> kerberos_ban |] ==> A=A' & Ts=Ts' & B=B' & X = X'"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   278
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   279
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   280
apply (erule kerberos_ban.induct) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   281
apply (frule_tac [7] Oops_parts_spies) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   282
apply (frule_tac [5] Kb3_msg_in_parts_spies, simp_all)  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   283
(*Kb2: it can't be a new key*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   284
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   285
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   286
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   287
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   288
(** Lemma: the session key sent in msg Kb2 would be EXPIRED
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   289
    if the spy could see it!
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   290
**)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   291
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   292
lemma lemma2 [rule_format (no_asm)]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   293
     "[| A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban |]            
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   294
  ==> Says Server A                                             
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   295
          (Crypt (shrK A) {|Number Ts, Agent B, Key K,          
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   296
                            Crypt (shrK B) {|Number Ts, Agent A, Key K|}|}) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   297
         \<in> set evs -->                                          
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   298
      Key K \<in> analz (spies evs) --> Expired Ts evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   299
apply (erule kerberos_ban.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   300
apply (frule_tac [7] Says_Server_message_form)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   301
apply (frule_tac [5] Says_S_message_form [THEN disjE])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   302
apply (simp_all (no_asm_simp) add: less_SucI analz_insert_eq analz_insert_freshK pushes)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   303
txt{*Fake*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   304
apply spy_analz
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   305
txt{*Kb2*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   306
apply (blast intro: parts_insertI less_SucI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   307
txt{*Kb3*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   308
apply (case_tac "Aa \<in> bad")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   309
 prefer 2 apply (blast dest: A_trusts_K_by_Kb2 unique_session_keys)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   310
apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   311
txt{*Oops: PROOF FAILED if addIs below*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   312
apply (blast dest: unique_session_keys intro!: less_SucI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   313
done
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   314
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   315
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   316
(** CONFIDENTIALITY for the SERVER:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   317
                     Spy does not see the keys sent in msg Kb2 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   318
                     as long as they have NOT EXPIRED
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   319
**)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   320
lemma Confidentiality_S:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   321
     "[| Says Server A                                            
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   322
          (Crypt K' {|Number T, Agent B, Key K, X|}) \<in> set evs;   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   323
         ~ Expired T evs;                                         
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   324
         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban                 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   325
      |] ==> Key K \<notin> analz (spies evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   326
apply (frule Says_Server_message_form, assumption)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   327
apply (blast intro: lemma2)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   328
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   329
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   330
(**** THE COUNTERPART OF CONFIDENTIALITY 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   331
      [|...; Expired Ts evs; ...|] ==> Key K \<in> analz (spies evs)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   332
      WOULD HOLD ONLY IF AN OOPS OCCURRED! ---> Nothing to prove!   ****)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   333
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   334
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   335
(** CONFIDENTIALITY for ALICE: **)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   336
(** Also A_trusts_K_by_Kb2 RS Confidentiality_S **)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   337
lemma Confidentiality_A:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   338
     "[| Crypt (shrK A) {|Number T, Agent B, Key K, X|} \<in> parts (spies evs); 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   339
         ~ Expired T evs;           
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   340
         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban                 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   341
      |] ==> Key K \<notin> analz (spies evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   342
apply (blast dest!: A_trusts_K_by_Kb2 Confidentiality_S)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   343
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   344
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   345
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   346
(** CONFIDENTIALITY for BOB: **)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   347
(** Also B_trusts_K_by_Kb3 RS Confidentiality_S **)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   348
lemma Confidentiality_B:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   349
     "[| Crypt (shrK B) {|Number Tk, Agent A, Key K|}  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   350
          \<in> parts (spies evs);               
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   351
        ~ Expired Tk evs;           
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   352
        A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban                 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   353
      |] ==> Key K \<notin> analz (spies evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   354
apply (blast dest!: B_trusts_K_by_Kb3 Confidentiality_S)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   355
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   356
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   357
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   358
lemma lemma_B [rule_format]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   359
     "[| B \<notin> bad;  evs \<in> kerberos_ban |]                         
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   360
      ==> Key K \<notin> analz (spies evs) -->                     
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   361
          Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   362
          \<in> set evs -->                                              
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   363
          Crypt K (Number Ta) \<in> parts (spies evs) -->         
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   364
          Says B A (Crypt K (Number Ta)) \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   365
apply (erule kerberos_ban.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   366
apply (frule_tac [7] Oops_parts_spies)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   367
apply (frule_tac [5] Says_S_message_form)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   368
apply (drule_tac [6] Kb3_msg_in_parts_spies, analz_mono_contra)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   369
apply (simp_all (no_asm_simp) add: all_conj_distrib)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   370
txt{*Fake*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   371
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   372
txt{*Kb2*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   373
apply (force dest: Crypt_imp_invKey_keysFor) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   374
txt{*Kb4*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   375
apply (blast dest: B_trusts_K_by_Kb3 unique_session_keys 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   376
                   Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   377
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   378
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   379
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   380
(*AUTHENTICATION OF B TO A*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   381
lemma Authentication_B:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   382
     "[| Crypt K (Number Ta) \<in> parts (spies evs);            
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   383
         Crypt (shrK A) {|Number Ts, Agent B, Key K, X|}     
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   384
         \<in> parts (spies evs);                                
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   385
         ~ Expired Ts evs;                                   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   386
         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban |]         
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   387
      ==> Says B A (Crypt K (Number Ta)) \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   388
by (blast dest!: A_trusts_K_by_Kb2
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   389
          intro!: lemma_B elim!: Confidentiality_S [THEN [2] rev_notE])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   390
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   391
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   392
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   393
lemma lemma_A [rule_format]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   394
     "[| A \<notin> bad; B \<notin> bad; evs \<in> kerberos_ban |] 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   395
      ==>           
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   396
         Key K \<notin> analz (spies evs) -->          
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   397
         Says Server A (Crypt (shrK A) {|Number Ts, Agent B, Key K, X|})   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   398
         \<in> set evs -->   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   399
          Crypt K {|Agent A, Number Ta|} \<in> parts (spies evs) --> 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   400
         Says A B {|X, Crypt K {|Agent A, Number Ta|}|}   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   401
             \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   402
apply (erule kerberos_ban.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   403
apply (frule_tac [7] Oops_parts_spies)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   404
apply (frule_tac [5] Says_S_message_form)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   405
apply (frule_tac [6] Kb3_msg_in_parts_spies, analz_mono_contra)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   406
apply (simp_all (no_asm_simp) add: all_conj_distrib)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   407
txt{*Fake*}
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
txt{*Kb2*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   410
apply (force dest: Crypt_imp_invKey_keysFor) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   411
txt{*Kb3*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   412
apply (blast dest: A_trusts_K_by_Kb2 unique_session_keys)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   413
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   414
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   415
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   416
(*AUTHENTICATION OF A TO B*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   417
lemma Authentication_A:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   418
     "[| Crypt K {|Agent A, Number Ta|} \<in> parts (spies evs);   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   419
         Crypt (shrK B) {|Number Ts, Agent A, Key K|}          
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   420
         \<in> parts (spies evs);                                  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   421
         ~ Expired Ts evs;                                     
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   422
         A \<notin> bad;  B \<notin> bad;  evs \<in> kerberos_ban |]           
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   423
      ==> Says A B {|Crypt (shrK B) {|Number Ts, Agent A, Key K|},      
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   424
                     Crypt K {|Agent A, Number Ta|}|} \<in> set evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   425
by (blast dest!: B_trusts_K_by_Kb3
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   426
          intro!: lemma_A 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13507
diff changeset
   427
          elim!: Confidentiality_S [THEN [2] rev_notE])
5053
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   428
75d20f367e94 New example Kerberos_BAN by G Bella
paulson
parents:
diff changeset
   429
end