src/HOL/Auth/Event.ML
author paulson
Tue, 03 Sep 1996 17:54:39 +0200
changeset 1942 6c9c1a42a869
parent 1933 8b24773de6db
child 2032 1bbf1bdcaf56
permissions -rw-r--r--
Renaming and simplification
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Message
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     2
    ID:         $Id$
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     5
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     6
Datatype of events;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     7
inductive relation "traces" for Needham-Schroeder (shared keys)
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
     8
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
     9
Rewrites should not refer to	 initState (Friend i)    -- not in normal form
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    10
*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    11
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    12
Addsimps [parts_cut_eq];
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    13
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    14
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    15
(*INSTALLED ON NAT.ML*)
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    16
goalw Nat.thy [le_def] "(m <= n) = (m < Suc n)";
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    17
by (rtac not_less_eq 1);
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    18
qed "le_eq_less_Suc";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    19
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    20
proof_timing:=true;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    21
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    22
(*FOR LIST.ML??*)
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    23
goal List.thy "x : set_of_list (x#xs)";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    24
by (Simp_tac 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    25
qed "set_of_list_I1";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    26
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    27
goal List.thy "!!x. xs = x#xs' ==> x : set_of_list xs";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    28
by (Asm_simp_tac 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    29
qed "set_of_list_eqI1";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    30
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    31
(*FUN.ML??  WE NEED A NOTION OF INVERSE IMAGE, OR GRAPH!!*)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    32
goal Set.thy "!!f. B <= range f = (B = f`` {x. f x: B})";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    33
by (fast_tac (!claset addEs [equalityE]) 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    34
val subset_range_iff = result();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    35
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    36
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    37
open Event;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    38
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    39
Addsimps [Un_insert_left, Un_insert_right];
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    40
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    41
(*By default only o_apply is built-in.  But in the presence of eta-expansion
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    42
  this means that some terms displayed as (f o g) will be rewritten, and others
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    43
  will not!*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    44
Addsimps [o_def];
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
    45
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    46
(*** Basic properties of shrK and newK ***)
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
    47
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    48
(* invKey (shrK A) = shrK A *)
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    49
bind_thm ("invKey_shrK", rewrite_rule [isSymKey_def] isSym_shrK);
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    50
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
    51
(* invKey (newK evs) = newK evs *)
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    52
bind_thm ("invKey_newK", rewrite_rule [isSymKey_def] isSym_newK);
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    53
Addsimps [invKey_shrK, invKey_newK];
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    54
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    55
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    56
(*New keys and nonces are fresh*)
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    57
val shrK_inject = inj_shrK RS injD;
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
    58
val newN_inject = inj_newN RS injD
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
    59
and newK_inject = inj_newK RS injD;
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    60
AddSEs [shrK_inject, newN_inject, newK_inject,
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    61
	fresh_newK RS notE, fresh_newN RS notE];
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    62
Addsimps [inj_shrK RS inj_eq, inj_newN RS inj_eq, inj_newK RS inj_eq];
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    63
Addsimps [fresh_newN, fresh_newK];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    64
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    65
goal thy "newK evs ~= shrK B";
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    66
by (subgoal_tac "newK evs = shrK B --> \
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
    67
\                Key (newK evs) : parts (initState B)" 1);
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    68
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    69
by (agent.induct_tac "B" 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    70
by (auto_tac (!claset addIs [range_eqI], !simpset));
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    71
qed "newK_neq_shrK";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    72
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
    73
Addsimps [newK_neq_shrK, newK_neq_shrK RS not_sym];
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    74
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    75
(*Good for talking about Server's initial state*)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    76
goal thy "!!H. H <= Key``E ==> parts H = H";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    77
by (Auto_tac ());
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    78
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    79
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    80
qed "parts_image_subset";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    81
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    82
bind_thm ("parts_image_Key", subset_refl RS parts_image_subset);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    83
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    84
goal thy "!!H. H <= Key``E ==> analz H = H";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    85
by (Auto_tac ());
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    86
be analz.induct 1;
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    87
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    88
qed "analz_image_subset";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    89
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    90
bind_thm ("analz_image_Key", subset_refl RS analz_image_subset);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
    91
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
    92
Addsimps [parts_image_Key, analz_image_Key];
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    93
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    94
goalw thy [keysFor_def] "keysFor (parts (initState C)) = {}";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    95
by (agent.induct_tac "C" 1);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
    96
by (auto_tac (!claset addIs [range_eqI] delrules partsEs, !simpset));
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    97
qed "keysFor_parts_initState";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    98
Addsimps [keysFor_parts_initState];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    99
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   100
goalw thy [keysFor_def] "keysFor (Key``E) = {}";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   101
by (Auto_tac ());
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   102
qed "keysFor_image_Key";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   103
Addsimps [keysFor_image_Key];
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   104
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   105
goal thy "shrK A ~: newK``E";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   106
by (agent.induct_tac "A" 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   107
by (Auto_tac ());
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   108
qed "shrK_notin_image_newK";
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   109
Addsimps [shrK_notin_image_newK];
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   110
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   111
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   112
(*Agents see their own shrKs!*)
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   113
goal thy "Key (shrK A) : analz (sees A evs)";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   114
by (list.induct_tac "evs" 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   115
by (asm_simp_tac (!simpset addsimps [impOfSubs(Un_upper2 RS analz_mono)]) 2);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   116
by (agent.induct_tac "A" 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   117
by (auto_tac (!claset addIs [range_eqI], !simpset));
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   118
qed "analz_own_shrK";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   119
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   120
bind_thm ("parts_own_shrK",
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   121
	  [analz_subset_parts, analz_own_shrK] MRS subsetD);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   122
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   123
Addsimps [analz_own_shrK, parts_own_shrK];
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   124
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   125
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   126
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   127
(**** Inductive relation "traces" -- basic properties ****)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   128
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   129
val mk_cases = traces.mk_cases (list.simps @ agent.simps @ event.simps);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   130
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   131
val Says_tracesE        = mk_cases "Says A B X # evs: traces";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   132
val Says_Server_tracesE = mk_cases "Says Server B X # evs: traces";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   133
val Says_Enemy_tracesE  = mk_cases "Says Enemy B X # evs: traces";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   134
val Says_to_Server_tracesE = mk_cases "Says A Server X # evs: traces";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   135
val Notes_tracesE       = mk_cases "Notes A X # evs: traces";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   136
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   137
(*The tail of a trace is a trace*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   138
goal thy "!!ev evs. ev#evs : traces ==> evs : traces";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   139
by (fast_tac (!claset addEs [mk_cases "ev#evs:traces"]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   140
qed "traces_ConsE";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   141
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   142
goal thy "!!ev evs. [| evs = ev#evs'; evs : traces |] ==> evs' : traces";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   143
by (fast_tac (!claset addEs [mk_cases "ev#evs:traces"]) 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   144
qed "traces_eq_ConsE";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   145
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   146
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   147
(** Specialized rewrite rules for (sees A (Says...#evs)) **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   148
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   149
goal thy "sees A (Says A B X # evs) = insert X (sees A evs)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   150
by (Simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   151
qed "sees_own";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   152
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   153
goal thy "!!A. Server ~= A ==> \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   154
\              sees Server (Says A B X # evs) = sees Server evs";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   155
by (Asm_simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   156
qed "sees_Server";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   157
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   158
goal thy "!!A. Friend i ~= A ==> \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   159
\              sees (Friend i) (Says A B X # evs) = sees (Friend i) evs";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   160
by (Asm_simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   161
qed "sees_Friend";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   162
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   163
goal thy "sees Enemy (Says A B X # evs) = insert X (sees Enemy evs)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   164
by (Simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   165
qed "sees_Enemy";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   166
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   167
goal thy "sees A (Says A' B X # evs) <= insert X (sees A evs)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   168
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   169
by (Fast_tac 1);
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   170
qed "sees_Says_subset_insert";
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   171
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   172
goal thy "sees A evs <= sees A (Says A' B X # evs)";
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   173
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   174
by (Fast_tac 1);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   175
qed "sees_subset_sees_Says";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   176
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   177
(*Pushing Unions into parts; one of the A's equals B, and thus sees Y*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   178
goal thy "(UN A. parts (sees A (Says B C Y # evs))) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   179
\         parts {Y} Un (UN A. parts (sees A evs))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   180
by (Step_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   181
be rev_mp 1;	(*for some reason, split_tac does not work on assumptions*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   182
val ss = (!simpset addsimps [parts_Un, sees_Cons] 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   183
	           setloop split_tac [expand_if]);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   184
by (ALLGOALS (fast_tac (!claset addss ss)));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   185
qed "UN_parts_sees_Says";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   186
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   187
goal thy "Says A B X : set_of_list evs --> X : sees Enemy evs";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   188
by (list.induct_tac "evs" 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   189
by (Auto_tac ());
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   190
qed_spec_mp "Says_imp_sees_Enemy";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   191
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   192
Addsimps [Says_imp_sees_Enemy];
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   193
AddIs    [Says_imp_sees_Enemy];
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   194
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   195
goal thy "initState C <= Key `` range shrK";
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   196
by (agent.induct_tac "C" 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   197
by (Auto_tac ());
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   198
qed "initState_subset";
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   199
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   200
goal thy "X : sees C evs --> \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   201
\          (EX A B. Says A B X : set_of_list evs) | \
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   202
\          (EX A. Notes A X : set_of_list evs) | \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   203
\          (EX A. X = Key (shrK A))";
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   204
by (list.induct_tac "evs" 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   205
by (ALLGOALS Asm_simp_tac);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   206
by (fast_tac (!claset addDs [impOfSubs initState_subset]) 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   207
br conjI 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   208
by (Fast_tac 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   209
by (event.induct_tac "a" 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   210
by (ALLGOALS (asm_simp_tac (!simpset addsimps [mem_if])));
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   211
by (ALLGOALS Fast_tac);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   212
qed_spec_mp "seesD";
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   213
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   214
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   215
Addsimps [UN_parts_sees_Says, sees_own, sees_Server, sees_Friend, sees_Enemy];
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   216
Delsimps [sees_Cons];	(**** NOTE REMOVAL -- laws above are cleaner ****)
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   217
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   218
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   219
(**** Inductive proofs about traces ****)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   220
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   221
(*The Enemy can see more than anybody else, except for their initial state*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   222
goal thy 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   223
 "!!evs. evs : traces ==> \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   224
\     sees A evs <= initState A Un sees Enemy evs";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   225
be traces.induct 1;
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   226
by (ALLGOALS (fast_tac (!claset addDs [sees_Says_subset_insert RS subsetD] 
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   227
			        addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   228
qed "sees_agent_subset_sees_Enemy";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   229
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   230
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   231
(*Nobody sends themselves messages*)
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   232
goal thy "!!evs. evs : traces ==> ALL A X. Says A A X ~: set_of_list evs";
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   233
be traces.induct 1;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   234
by (Auto_tac());
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   235
qed_spec_mp "not_Says_to_self";
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   236
Addsimps [not_Says_to_self];
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   237
AddSEs   [not_Says_to_self RSN (2, rev_notE)];
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   238
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   239
goal thy "!!evs. evs : traces ==> Notes A X ~: set_of_list evs";
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   240
be traces.induct 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   241
by (Auto_tac());
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   242
qed "not_Notes";
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   243
Addsimps [not_Notes];
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   244
AddSEs   [not_Notes RSN (2, rev_notE)];
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   245
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   246
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   247
goal thy "!!evs. (Says S A (Crypt {|N, B, K, X|} KA)) : set_of_list evs ==> \
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   248
\                X : parts (sees Enemy evs)";
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   249
by (fast_tac (!claset addSEs partsEs
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   250
	              addSDs [Says_imp_sees_Enemy RS parts.Inj]) 1);
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   251
qed "NS3_msg_in_parts_sees_Enemy";
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   252
			      
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   253
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   254
(*** Server keys are not betrayed ***)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   255
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   256
(*Enemy never sees another agent's server key!*)
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   257
goal thy 
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   258
 "!!evs. [| evs : traces; A ~= Enemy |] ==> \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   259
\        Key (shrK A) ~: parts (sees Enemy evs)";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   260
be traces.induct 1;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   261
bd NS3_msg_in_parts_sees_Enemy 5;
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   262
by (Auto_tac());
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   263
(*Deals with Fake message*)
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   264
by (best_tac (!claset addDs [impOfSubs analz_subset_parts,
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   265
			     impOfSubs synth_analz_parts_insert_subset_Un]) 1);
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   266
qed "Enemy_not_see_shrK";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   267
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   268
bind_thm ("Enemy_not_analz_shrK",
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   269
	  [analz_subset_parts, Enemy_not_see_shrK] MRS contra_subsetD);
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   270
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   271
Addsimps [Enemy_not_see_shrK, 
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   272
	  not_sym RSN (2, Enemy_not_see_shrK), 
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   273
	  Enemy_not_analz_shrK, 
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   274
	  not_sym RSN (2, Enemy_not_analz_shrK)];
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   275
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   276
(*We go to some trouble to preserve R in the 3rd subgoal*)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   277
val major::prems = 
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   278
goal thy  "[| Key (shrK A) : parts (sees Enemy evs);    \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   279
\             evs : traces;                                  \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   280
\             A=Enemy ==> R                                  \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   281
\           |] ==> R";
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   282
br ccontr 1;
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   283
br ([major, Enemy_not_see_shrK] MRS rev_notE) 1;
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   284
by (swap_res_tac prems 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   285
by (ALLGOALS (fast_tac (!claset addIs prems)));
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   286
qed "Enemy_see_shrK_E";
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   287
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   288
bind_thm ("Enemy_analz_shrK_E", 
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   289
	  analz_subset_parts RS subsetD RS Enemy_see_shrK_E);
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   290
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   291
(*Classical reasoner doesn't need the not_sym versions (with swapped ~=) *)
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   292
AddSEs [Enemy_see_shrK_E, Enemy_analz_shrK_E];
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   293
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   294
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   295
(*No Friend will ever see another agent's server key 
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   296
  (excluding the Enemy, who might transmit his).
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   297
  The Server, of course, knows all server keys.*)
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   298
goal thy 
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   299
 "!!evs. [| evs : traces; A ~= Enemy;  A ~= Friend j |] ==> \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   300
\        Key (shrK A) ~: parts (sees (Friend j) evs)";
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   301
br (sees_agent_subset_sees_Enemy RS parts_mono RS contra_subsetD) 1;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   302
by (ALLGOALS Asm_simp_tac);
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   303
qed "Friend_not_see_shrK";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   304
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   305
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   306
(*Not for Addsimps -- it can cause goals to blow up!*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   307
goal thy  
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   308
 "!!evs. evs : traces ==>                                  \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   309
\        (Key (shrK A) \
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   310
\           : analz (insert (Key (shrK B)) (sees Enemy evs))) =  \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   311
\        (A=B | A=Enemy)";
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   312
by (best_tac (!claset addDs [impOfSubs analz_subset_parts]
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   313
		      addIs [impOfSubs (subset_insertI RS analz_mono)]
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   314
	              addss (!simpset)) 1);
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   315
qed "shrK_mem_analz";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   316
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   317
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   318
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   319
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   320
(*** Future keys can't be seen or used! ***)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   321
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   322
(*Nobody can have SEEN keys that will be generated in the future.
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   323
  This has to be proved anew for each protocol description,
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   324
  but should go by similar reasoning every time.  Hardest case is the
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   325
  standard Fake rule.  
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   326
      The length comparison, and Union over C, are essential for the 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   327
  induction! *)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   328
goal thy "!!evs. evs : traces ==> \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   329
\                length evs <= length evs' --> \
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   330
\                          Key (newK evs') ~: (UN C. parts (sees C evs))";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   331
be traces.induct 1;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   332
bd NS3_msg_in_parts_sees_Enemy 5;
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   333
(*auto_tac does not work here, as it performs safe_tac first*)
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   334
by (ALLGOALS Asm_simp_tac);
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   335
by (ALLGOALS (best_tac (!claset addDs [impOfSubs analz_subset_parts,
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   336
				       impOfSubs parts_insert_subset_Un,
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   337
				       Suc_leD]
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   338
			        addss (!simpset))));
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   339
val lemma = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   340
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   341
(*Variant needed for the main theorem below*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   342
goal thy 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   343
 "!!evs. [| evs : traces;  length evs <= length evs' |] ==> \
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   344
\        Key (newK evs') ~: parts (sees C evs)";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   345
by (fast_tac (!claset addDs [lemma]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   346
qed "new_keys_not_seen";
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   347
Addsimps [new_keys_not_seen];
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   348
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   349
(*Another variant: old messages must contain old keys!*)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   350
goal thy 
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   351
 "!!evs. [| Says A B X : set_of_list evs;  \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   352
\           Key (newK evt) : parts {X};    \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   353
\           evs : traces                 \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   354
\        |] ==> length evt < length evs";
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   355
br ccontr 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   356
by (fast_tac (!claset addSDs [new_keys_not_seen, Says_imp_sees_Enemy]
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   357
	              addIs [impOfSubs parts_mono, leI]) 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   358
qed "Says_imp_old_keys";
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   359
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   360
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   361
goal thy "!!K. newK evs = invKey K ==> newK evs = K";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   362
br (invKey_eq RS iffD1) 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   363
by (Simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   364
val newK_invKey = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   365
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   366
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   367
val keysFor_parts_mono = parts_mono RS keysFor_mono;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   368
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   369
(*Nobody can have USED keys that will be generated in the future.
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   370
  ...very like new_keys_not_seen*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   371
goal thy "!!evs. evs : traces ==> \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   372
\                length evs <= length evs' --> \
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   373
\                newK evs' ~: keysFor (UN C. parts (sees C evs))";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   374
be traces.induct 1;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   375
bd NS3_msg_in_parts_sees_Enemy 5;
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   376
by (ALLGOALS Asm_simp_tac);
1930
cdf9bcd53749 Working version of NS, messages 1-4!
paulson
parents: 1929
diff changeset
   377
(*NS1 and NS2*)
cdf9bcd53749 Working version of NS, messages 1-4!
paulson
parents: 1929
diff changeset
   378
map (by o fast_tac (!claset addDs [Suc_leD] addss (!simpset))) [3,2];
cdf9bcd53749 Working version of NS, messages 1-4!
paulson
parents: 1929
diff changeset
   379
(*Fake and NS3*)
cdf9bcd53749 Working version of NS, messages 1-4!
paulson
parents: 1929
diff changeset
   380
map (by o best_tac
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   381
     (!claset addSDs [newK_invKey]
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   382
	      addDs [impOfSubs (analz_subset_parts RS keysFor_mono),
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   383
		     impOfSubs (parts_insert_subset_Un RS keysFor_mono),
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   384
		     Suc_leD]
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   385
	      addEs [new_keys_not_seen RS not_parts_not_analz RSN (2,rev_notE)]
1930
cdf9bcd53749 Working version of NS, messages 1-4!
paulson
parents: 1929
diff changeset
   386
	      addss (!simpset)))
cdf9bcd53749 Working version of NS, messages 1-4!
paulson
parents: 1929
diff changeset
   387
    [2,1];
1933
8b24773de6db Addition of message NS5
paulson
parents: 1930
diff changeset
   388
(*NS4 and NS5: nonce exchange*)
8b24773de6db Addition of message NS5
paulson
parents: 1930
diff changeset
   389
by (ALLGOALS (deepen_tac (!claset addSDs [newK_invKey, Says_imp_old_keys]
8b24773de6db Addition of message NS5
paulson
parents: 1930
diff changeset
   390
	                          addIs  [less_SucI, impOfSubs keysFor_mono]
8b24773de6db Addition of message NS5
paulson
parents: 1930
diff changeset
   391
		                  addss (!simpset addsimps [le_def])) 0));
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   392
val lemma = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   393
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   394
goal thy 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   395
 "!!evs. [| evs : traces;  length evs <= length evs' |] ==> \
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   396
\        newK evs' ~: keysFor (parts (sees C evs))";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   397
by (fast_tac (!claset addSDs [lemma] addss (!simpset)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   398
qed "new_keys_not_used";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   399
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   400
bind_thm ("new_keys_not_analzd",
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   401
	  [analz_subset_parts RS keysFor_mono,
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   402
	   new_keys_not_used] MRS contra_subsetD);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   403
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   404
Addsimps [new_keys_not_used, new_keys_not_analzd];
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   405
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   406
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   407
(** Rewrites to push in Key and Crypt messages, so that other messages can
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   408
    be pulled out using the analz_insert rules **)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   409
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   410
fun insComm x y = read_instantiate_sg (sign_of thy) [("x",x), ("y",y)] 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   411
                      insert_commute;
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   412
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   413
val pushKeys = map (insComm "Key ?K") 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   414
                  ["Agent ?C", "Nonce ?N", "MPair ?X ?Y", "Crypt ?X ?K'"];
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   415
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   416
val pushCrypts = map (insComm "Crypt ?X ?K") 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   417
                    ["Agent ?C", "Nonce ?N", "MPair ?X' ?Y"];
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   418
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   419
(*Cannot be added with Addsimps -- we don't always want to re-order messages*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   420
val pushes = pushKeys@pushCrypts;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   421
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   422
val pushKey_newK = insComm "Key (newK ?evs)"  "Key (shrK ?C)";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   423
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   424
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   425
(** Lemmas concerning the form of items passed in messages **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   426
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   427
(*Describes the form *and age* of K, and the form of X,
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   428
  when the following message is sent*)
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   429
goal thy 
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   430
 "!!evs. [| Says Server A (Crypt {|N, Agent B, K, X|} K') : set_of_list evs; \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   431
\           evs : traces    \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   432
\        |] ==> (EX evt:traces. \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   433
\                         K = Key(newK evt) & \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   434
\                         X = (Crypt {|K, Agent A|} (shrK B)) & \
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   435
\                         K' = shrK A & \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   436
\                         length evt < length evs)";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   437
be rev_mp 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   438
be traces.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   439
by (ALLGOALS (fast_tac (!claset addIs [less_SucI] addss (!simpset))));
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   440
qed "Says_Server_message_form";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   441
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   442
(* c ~: keysFor (parts H) ==> c ~: keysFor (analz H) *)
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   443
bind_thm ("not_parts_not_keysFor_analz", 
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   444
	  analz_subset_parts RS keysFor_mono RS contra_subsetD);
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   445
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   446
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   447
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   448
(*USELESS for NS3, case 1, as the ind hyp says the same thing!
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   449
goal thy 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   450
 "!!evs. [| evs = Says Server (Friend i) \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   451
\                 (Crypt {|N, Agent(Friend j), K, X|} K') # evs';  \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   452
\           evs : traces;  i~=k                                    \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   453
\        |] ==>                                                    \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   454
\     K ~: analz (insert (Key (shrK (Friend k))) (sees Enemy evs))";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   455
be rev_mp 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   456
be traces.induct 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   457
by (ALLGOALS (asm_full_simp_tac (!simpset addsimps pushes)));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   458
by (Step_tac 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   459
by (asm_full_simp_tac (!simpset addsimps[analz_subset_parts RS contra_subsetD]) 1);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   460
val Enemy_not_see_encrypted_key_lemma = result();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   461
*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   462
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   463
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   464
(*Describes the form of X when the following message is sent*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   465
goal thy
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   466
 "!!evs. evs : traces ==>                             \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   467
\        ALL A NA B K X.                            \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   468
\            (Crypt {|Nonce NA, Agent B, Key K, X|} (shrK A)) \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   469
\            : parts (sees Enemy evs) & A ~= Enemy  -->   \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   470
\          (EX evt:traces. K = newK evt & \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   471
\                          X = (Crypt {|Key K, Agent A|} (shrK B)))";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   472
be traces.induct 1;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   473
bd NS3_msg_in_parts_sees_Enemy 5;
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   474
by (Step_tac 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   475
by (ALLGOALS Asm_full_simp_tac);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   476
(*Remaining cases are Fake and NS2*)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   477
by (fast_tac (!claset addSDs [spec]) 2);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   478
(*Now for the Fake case*)
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   479
by (best_tac (!claset addDs [impOfSubs analz_subset_parts,
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   480
			     impOfSubs synth_analz_parts_insert_subset_Un]
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   481
	              addss (!simpset)) 1);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   482
qed_spec_mp "encrypted_form";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   483
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   484
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   485
(*For eliminating the A ~= Enemy condition from the previous result*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   486
goal thy 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   487
 "!!evs. evs : traces ==>                             \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   488
\        ALL S A NA B K X.                            \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   489
\            Says S A (Crypt {|Nonce NA, Agent B, Key K, X|} (shrK A)) \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   490
\            : set_of_list evs  -->   \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   491
\        S = Server | S = Enemy";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   492
be traces.induct 1;
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   493
by (ALLGOALS Asm_simp_tac);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   494
(*We are left with NS3*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   495
by (subgoal_tac "S = Server | S = Enemy" 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   496
(*First justify this assumption!*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   497
by (fast_tac (!claset addSEs [allE, mp] addss (!simpset)) 2);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   498
by (Step_tac 1);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   499
bd Says_Server_message_form 1;
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   500
by (ALLGOALS Full_simp_tac);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   501
(*Final case.  Clear out needless quantifiers to speed the following step*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   502
by (eres_inst_tac [("V","ALL x. ?P(x)")] thin_rl 1);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   503
bd encrypted_form 1;
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   504
br (parts.Inj RS conjI) 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   505
auto();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   506
qed_spec_mp "Server_or_Enemy";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   507
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   508
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   509
(*Describes the form of X when the following message is sent;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   510
  use Says_Server_message_form if applicable*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   511
goal thy 
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   512
 "!!evs. [| Says S A (Crypt {|Nonce NA, Agent B, Key K, X|} (shrK A)) \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   513
\            : set_of_list evs;                              \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   514
\           evs : traces               \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   515
\        |] ==> (EX evt:traces. K = newK evt & length evt < length evs & \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   516
\                               X = (Crypt {|Key K, Agent A|} (shrK B)))";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   517
by (forward_tac [Server_or_Enemy] 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   518
ba 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   519
by (Step_tac 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   520
by (fast_tac (!claset addSDs [Says_Server_message_form] addss (!simpset)) 1);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   521
by (forward_tac [encrypted_form] 1);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   522
br (parts.Inj RS conjI) 1;
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   523
by (auto_tac (!claset addIs [Says_imp_old_keys], !simpset));
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   524
qed "Says_S_message_form";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   525
1930
cdf9bcd53749 Working version of NS, messages 1-4!
paulson
parents: 1929
diff changeset
   526
(*Currently unused.  Needed only to reason about the VERY LAST message.*)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   527
goal thy 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   528
 "!!evs. [| evs = Says S A (Crypt {|Nonce NA, Agent B, Key K, X|}   \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   529
\                           (shrK A)) # evs';                  \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   530
\           evs : traces                                           \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   531
\        |] ==> (EX evt:traces. evs' : traces & length evt < length evs & \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   532
\                               K = newK evt & \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   533
\                               X = (Crypt {|Key K, Agent A|} (shrK B)))";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   534
by (forward_tac [traces_eq_ConsE] 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   535
by (dtac (set_of_list_eqI1 RS Says_S_message_form) 2);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   536
by (Auto_tac());	
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   537
qed "Says_S_message_form_eq";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   538
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   539
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   540
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   541
(****
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   542
 The following is to prove theorems of the form
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   543
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   544
          Key K : analz (insert (Key (newK evt)) 
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   545
	                   (insert (Key (shrK C)) (sees Enemy evs))) ==>
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   546
          Key K : analz (insert (Key (shrK C)) (sees Enemy evs))
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   547
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   548
 A more general formula must be proved inductively.
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   549
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   550
****)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   551
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   552
1933
8b24773de6db Addition of message NS5
paulson
parents: 1930
diff changeset
   553
(*NOT useful in this form, but it says that session keys are not used
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   554
  to encrypt messages containing other keys, in the actual protocol.
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   555
  We require that agents should behave like this subsequently also.*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   556
goal thy 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   557
 "!!evs. evs : traces ==> \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   558
\        (Crypt X (newK evt)) : parts (sees Enemy evs) & \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   559
\        Key K : parts {X} --> Key K : parts (sees Enemy evs)";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   560
be traces.induct 1;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   561
bd NS3_msg_in_parts_sees_Enemy 5;
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   562
by (ALLGOALS (asm_simp_tac (!simpset addsimps pushes)));
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   563
(*Deals with Faked messages*)
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   564
by (best_tac (!claset addSEs partsEs
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   565
		      addDs [impOfSubs analz_subset_parts,
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   566
                             impOfSubs parts_insert_subset_Un]
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   567
                      addss (!simpset)) 1);
1933
8b24773de6db Addition of message NS5
paulson
parents: 1930
diff changeset
   568
(*NS4 and NS5*)
8b24773de6db Addition of message NS5
paulson
parents: 1930
diff changeset
   569
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   570
result();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   571
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   572
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   573
(** Specialized rewriting for this proof **)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   574
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   575
Delsimps [image_insert];
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   576
Addsimps [image_insert RS sym];
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   577
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   578
goal thy "insert (Key (newK x)) (sees A evs) = \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   579
\         Key `` (newK``{x}) Un (sees A evs)";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   580
by (Fast_tac 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   581
val insert_Key_singleton = result();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   582
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   583
goal thy "insert (Key (f x)) (Key``(f``E) Un C) = \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   584
\         Key `` (f `` (insert x E)) Un C";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   585
by (Fast_tac 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   586
val insert_Key_image = result();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   587
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   588
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   589
(** Session keys are not used to encrypt other session keys **)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   590
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   591
goal thy  
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   592
 "!!evs. evs : traces ==> \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   593
\  ALL K E. (Key K : analz (insert (Key (shrK C)) \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   594
\                             (Key``(newK``E) Un (sees Enemy evs)))) = \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   595
\           (K : newK``E |  \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   596
\            Key K : analz (insert (Key (shrK C)) \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   597
\                             (sees Enemy evs)))";
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   598
be traces.induct 1;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   599
by (forward_tac [Says_S_message_form] 5 THEN assume_tac 5);	
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   600
by (REPEAT ((eresolve_tac [bexE, conjE] ORELSE' hyp_subst_tac) 5));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   601
by (ALLGOALS 
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   602
    (asm_simp_tac 
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   603
     (!simpset addsimps ([insert_Key_singleton, insert_Key_image, pushKey_newK]
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   604
			 @ pushes)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   605
               setloop split_tac [expand_if])));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   606
(*Cases NS2 and NS3!!  Simple, thanks to auto case splits*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   607
by (REPEAT (Fast_tac 3));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   608
(*Base case*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   609
by (fast_tac (!claset addIs [image_eqI] addss (!simpset)) 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   610
(** LEVEL 7 **)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   611
(*Fake case*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   612
by (REPEAT (Safe_step_tac 1));
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   613
by (fast_tac (!claset addSEs [impOfSubs analz_mono]) 2);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   614
by (subgoal_tac 
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   615
    "Key K : analz \
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   616
\             (synth \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   617
\              (analz (insert (Key (shrK C)) \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   618
\                        (Key``(newK``E) Un (sees Enemy evsa)))))" 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   619
(*First, justify this subgoal*)
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   620
(*Discard formulae for better speed*)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   621
by (eres_inst_tac [("V","ALL S.?P(S)")] thin_rl 2);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   622
by (eres_inst_tac [("V","?Q ~: ?QQ")] thin_rl 2);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   623
by (best_tac (!claset addIs [impOfSubs (analz_mono RS synth_mono)]
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   624
                      addSEs [impOfSubs analz_mono]) 2);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   625
by (Asm_full_simp_tac 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   626
by (deepen_tac (!claset addIs [impOfSubs analz_mono]) 0 1);
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   627
qed_spec_mp "analz_image_newK";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   628
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   629
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   630
goal thy  
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   631
 "!!evs. evs : traces ==>                                  \
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   632
\        Key K : analz (insert (Key (newK evt))            \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   633
\                         (insert (Key (shrK C))      \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   634
\                          (sees Enemy evs))) =            \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   635
\             (K = newK evt |                              \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   636
\              Key K : analz (insert (Key (shrK C))   \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   637
\                               (sees Enemy evs)))";
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   638
by (asm_simp_tac (HOL_ss addsimps [pushKey_newK, analz_image_newK, 
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   639
				   insert_Key_singleton]) 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   640
by (Fast_tac 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   641
qed "analz_insert_Key_newK";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   642
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   643
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   644
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   645
(*This says that the Key, K, uniquely identifies the message.
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   646
    But if C=Enemy then he could send all sorts of nonsense.*)
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   647
goal thy 
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   648
 "!!evs. evs : traces ==>                      \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   649
\      EX X'. ALL C S A Y N B X.               \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   650
\         C ~= Enemy -->                       \
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   651
\         Says S A Y : set_of_list evs -->     \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   652
\         ((Crypt {|N, Agent B, Key K, X|} (shrK C)) : parts{Y} --> \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   653
\       (X = X'))";
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   654
be traces.induct 1;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   655
by (forward_tac [Says_S_message_form] 5 THEN assume_tac 5);	
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   656
by (ALLGOALS 
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   657
    (asm_simp_tac (!simpset addsimps [all_conj_distrib, imp_conj_distrib])));
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   658
(*NS2: Case split propagates some context to other subgoal...*)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   659
by (excluded_middle_tac "K = newK evsa" 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   660
by (Asm_simp_tac 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   661
(*...we assume X is a very new message, and handle this case by contradiction*)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   662
by (fast_tac (!claset addIs [impOfSubs (subset_insertI RS parts_mono)]
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   663
		      addSEs partsEs
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   664
		      addEs [Says_imp_old_keys RS less_irrefl]
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   665
	              addss (!simpset)) 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   666
(*NS3: No relevant messages*)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   667
by (fast_tac (!claset addSEs [exI] addss (!simpset)) 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   668
(*Fake*)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   669
by (Step_tac 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   670
br exI 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   671
br conjI 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   672
ba 2;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   673
by (Step_tac 1);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   674
(** LEVEL 12 **)
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   675
by (subgoal_tac "Crypt {|N, Agent Ba, Key K, Xa|} (shrK C) \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   676
\                  : parts (sees Enemy evsa)" 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   677
by (eres_inst_tac [("V","ALL S.?P(S)")] thin_rl 2);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   678
by (best_tac (!claset addSEs [impOfSubs analz_subset_parts]
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   679
	              addDs [impOfSubs parts_insert_subset_Un]
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   680
                      addss (!simpset)) 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   681
by (eres_inst_tac [("V","?aa : parts {X}")] thin_rl 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   682
bd parts_singleton 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   683
by (Step_tac 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   684
bd seesD 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   685
by (Step_tac 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   686
by (Full_simp_tac 2);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   687
by (fast_tac (!claset addSDs [spec]) 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   688
val lemma = result();
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   689
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   690
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   691
(*In messages of this form, the session key uniquely identifies the rest*)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   692
goal thy 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   693
 "!!evs. [| Says S A          \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   694
\             (Crypt {|N, Agent B, Key K, X|} (shrK C))     \
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   695
\                  : set_of_list evs; \
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   696
 \          Says S' A'                                         \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   697
\             (Crypt {|N', Agent B', Key K, X'|} (shrK C')) \
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   698
\                  : set_of_list evs;                         \
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   699
\           evs : traces;  C ~= Enemy;  C' ~= Enemy    |] ==> X = X'";
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   700
bd lemma 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   701
be exE 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   702
by (forw_inst_tac [("psi", "ALL C.?P(C)")] asm_rl 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   703
by (Fast_tac 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   704
qed "unique_session_keys";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   705
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   706
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   707
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   708
(*Crucial security property: Enemy does not see the keys sent in msg NS2
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   709
   -- even if another key is compromised*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   710
goal thy 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   711
 "!!evs. [| Says Server (Friend i) \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   712
\            (Crypt {|N, Agent(Friend j), K, X|} K') : set_of_list evs;  \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   713
\           evs : traces;  Friend i ~= C;  Friend j ~= C              \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   714
\        |] ==>                                                       \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   715
\     K ~: analz (insert (Key (shrK C)) (sees Enemy evs))";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   716
be rev_mp 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   717
be traces.induct 1;
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   718
by (ALLGOALS (asm_simp_tac (!simpset addsimps pushes)));
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   719
(*Next 3 steps infer that K has the form "Key (newK evs'" ... *)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   720
by (REPEAT_FIRST (resolve_tac [conjI, impI]));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   721
by (TRYALL (forward_tac [Says_Server_message_form] THEN' assume_tac));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   722
by (REPEAT_FIRST (eresolve_tac [bexE, conjE] ORELSE' hyp_subst_tac));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   723
by (ALLGOALS 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   724
    (asm_full_simp_tac 
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   725
     (!simpset addsimps ([analz_subset_parts RS contra_subsetD,
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   726
			  analz_insert_Key_newK] @ pushes)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   727
               setloop split_tac [expand_if])));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   728
(*NS2*)
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   729
by (fast_tac (!claset addSEs [less_irrefl]) 2);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   730
(** LEVEL 8 **)
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   731
(*Now for the Fake case*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   732
br notI 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   733
by (subgoal_tac 
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   734
    "Key (newK evt) : \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   735
\    analz (synth (analz (insert (Key (shrK C)) \
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   736
\                                  (sees Enemy evsa))))" 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   737
be (impOfSubs analz_mono) 2;
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   738
by (deepen_tac (!claset addIs [analz_mono RS synth_mono RSN (2,rev_subsetD),
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   739
			       impOfSubs synth_increasing,
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   740
			       impOfSubs analz_increasing]) 0 2);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   741
(*Proves the Fake goal*)
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   742
by (fast_tac (!claset addss (!simpset)) 1);
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   743
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   744
(**LEVEL 13**)
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   745
(*NS3: that message from the Server was sent earlier*)
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   746
by (mp_tac 1);
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   747
by (forward_tac [Says_S_message_form] 1 THEN assume_tac 1);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   748
by (REPEAT_FIRST (eresolve_tac [bexE, conjE] ORELSE' hyp_subst_tac));
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   749
by (asm_full_simp_tac
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   750
    (!simpset addsimps (mem_if::analz_insert_Key_newK::pushes)) 1);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   751
by (Step_tac 1);
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   752
(**LEVEL 18 **)
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   753
bd unique_session_keys 1;
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   754
by (REPEAT_FIRST assume_tac);
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   755
by (ALLGOALS Full_simp_tac);
1929
f0839bab4b00 Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents: 1913
diff changeset
   756
by (Step_tac 1);
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   757
by (asm_full_simp_tac (!simpset addsimps [shrK_mem_analz]) 1);
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   758
qed "Enemy_not_see_encrypted_key";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   759
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   760
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   761
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   762
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   763
XXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXX;
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   764
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   765
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   766
goals_limit:=4;
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   767
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   768
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   769
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   770
goal thy 
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   771
 "!!evs. [| Says Server (Friend i) \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   772
\             (Crypt {|N, Agent B, K|} K') : set_of_list evs;  \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   773
\           evs : traces;  i~=j    \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   774
\         |] ==> K ~: analz (sees (Friend j) evs)";
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   775
br (sees_agent_subset_sees_Enemy RS analz_mono RS contra_subsetD) 1;
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   776
by (ALLGOALS (asm_simp_tac (!simpset addsimps [Enemy_not_see_encrypted_key])));
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   777
qed "Friend_not_see_encrypted_key";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   778
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   779
goal thy 
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   780
 "!!evs. [| Says Server (Friend i) \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   781
\             (Crypt {|N, Agent B, K|} K') : set_of_list evs;  \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   782
\           A ~: {Friend i, Server};  \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   783
\           evs : traces    \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   784
\        |] ==>  K ~: analz (sees A evs)";
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   785
by (eres_inst_tac [("P", "A~:?X")] rev_mp 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   786
by (agent.induct_tac "A" 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   787
by (ALLGOALS Simp_tac);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   788
by (asm_simp_tac (!simpset addsimps [eq_sym_conv, 
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   789
				     Friend_not_see_encrypted_key]) 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   790
br ([analz_mono, Enemy_not_see_encrypted_key] MRS contra_subsetD) 1;
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   791
(*  hyp_subst_tac would deletes the equality assumption... *)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   792
by (ALLGOALS (rtac n_not_Suc_n ORELSE' Fast_tac));
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   793
qed "Agent_not_see_encrypted_key";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   794
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   795
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   796
(*Neatly packaged, perhaps in a useless form*)
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   797
goalw thy [knownBy_def]
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   798
 "!!evs. [| Says Server (Friend i) \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   799
\             (Crypt {|N, Agent B, K|} K') : set_of_list evs;  \
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   800
\           evs : traces    \
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   801
\        |] ==>  knownBy evs K <= {Friend i, Server}";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   802
1893
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   803
by (Step_tac 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   804
br ccontr 1;
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   805
by (fast_tac (!claset addDs [Agent_not_see_encrypted_key]) 1);
fa58f4a06f21 Works up to main theorem, then XXX...X
paulson
parents: 1885
diff changeset
   806
qed "knownBy_encrypted_key";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   807
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   808
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   809
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   810
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   811
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   812
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   813
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   814
push_proof();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   815
goal thy 
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   816
 "!!evs. [| evs = Says S (Friend i) \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   817
\                 (Crypt {|N, Agent(Friend j), K, X|} K') # evs';  \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   818
\           evs : traces;  i~=k                                    \
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   819
\        |] ==>                                                    \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   820
\     K ~: analz (insert (Key (shrK (Friend k))) (sees Enemy evs))";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   821
be rev_mp 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   822
be traces.induct 1;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   823
by (ALLGOALS (asm_full_simp_tac (!simpset addsimps pushes)));
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   824
by (Step_tac 1);
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   825
by (asm_full_simp_tac (!simpset addsimps[analz_subset_parts RS contra_subsetD]) 1);
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   826
val Enemy_not_see_encrypted_key_lemma = result();
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   827
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   828
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   829
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   830
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   831
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   832
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   833
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   834
(*Precisely formulated as needed below.  Perhaps redundant, as simplification
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   835
  with the aid of  analz_subset_parts RS contra_subsetD  might do the
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   836
  same thing.*)
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   837
goal thy 
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   838
 "!!evs. [| evs : traces; A ~= Enemy;  A ~= Friend j |] ==> \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   839
\        Key (shrK A) ~:                               \
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   840
\          analz (insert (Key (shrK (Friend j))) (sees Enemy evs))";
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   841
br (analz_subset_parts RS contra_subsetD) 1;
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   842
by (Asm_simp_tac 1);
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   843
qed "insert_not_analz_shrK";
1885
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   844
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   845
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   846
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   847
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   848
XXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXXX;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   849
proof_timing:=true;
a18a6dc14f76 Auth proofs work up to the XXX...
paulson
parents: 1852
diff changeset
   850
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   851
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   853
YYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYY;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   854
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   855
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   856
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   857
(*If a message is sent, encrypted with a Friend's server key, then either
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   858
  that Friend or the Server originally sent it.*)
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   859
goal thy 
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   860
 "!!evs. evs : traces ==>                             \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   861
\    ALL A B X i. Says A B (Crypt X (shrK (Friend i))) \
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   862
\        : set_of_list evs  -->   \
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   863
\    (EX B'. Says Server B' (Crypt X (shrK (Friend i))) : set_of_list evs | \
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   864
\            Says (Friend i) B' (Crypt X (shrK (Friend i))) : set_of_list evs)";
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   865
be traces.induct 1;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   866
by (Step_tac 1);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   867
by (ALLGOALS Asm_full_simp_tac);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   868
(*Remaining cases are Fake, NS2 and NS3*)
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   869
by (Fast_tac 2);	(*Proves the NS2 case*)
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   870
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   871
by (REPEAT (dtac spec 2));
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   872
fe conjE;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   873
bd mp 2;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   874
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   875
by (REPEAT (resolve_tac [refl, conjI] 2));
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   876
by (ALLGOALS Asm_full_simp_tac);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   877
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   878
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   879
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   880
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   881
by (full_simp_tac (!simpset addsimps [all_conj_distrib]) 2);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   882
be conjE 2;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   883
by ((dtac spec THEN' dtac spec THEN' dtac spec THEN' dtac spec)2);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   884
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   885
(*The NS3 case needs the induction hypothesis twice!
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   886
    One application is to the X component of the most recent message.*)
1942
6c9c1a42a869 Renaming and simplification
paulson
parents: 1933
diff changeset
   887
by (subgoal_tac "? evs'. X = Crypt {|Key (newK evs'), Agent A|} (shrK B)" 2);
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   888
by (Fast_tac 3);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   889
by (full_simp_tac (!simpset addsimps [all_conj_distrib]) 2);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   890
be conjE 2;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   891
(*DELETE the first quantified formula: it's now useless*)
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   892
by (eres_inst_tac [("V","ALL S.?P(S)")] thin_rl 2);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   893
by (fast_tac (!claset addss (!simpset)) 2);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   894
(*Now for the Fake case*)
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   895
be disjE 1;
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   896
(*The subcase of Fake, where the message in question is NOT the most recent*)
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   897
by (Best_tac 2);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   898
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   899
by (REPEAT_FIRST (etac conjE ORELSE' hyp_subst_tac));
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   900
be Crypt_synth 1;
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   901
be Key_synth 2;
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   902
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1893
diff changeset
   903
(*Split up the possibilities of that message being synthd...*)
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   904
by (Step_tac 1);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   905
by (Best_tac 6);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   906
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   907
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   908
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   909
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   910
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   911
(*NEEDED??*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   912
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   913
goal thy "!!A. X ~= Y ==> (X : sees A (Says B C Y # evs)) = (X : sees A evs)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   914
by (asm_simp_tac (!simpset addsimps [sees_Cons] 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   915
                           setloop split_tac [expand_if]) 1); 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   916
qed "in_sees_Says";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   917
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   918
goal thy "!!A. X ~: parts {Y} ==> \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   919
\              (X : parts (sees A (Says B C Y # evs))) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   920
\              (X : parts (sees A evs))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   921
by (asm_simp_tac (!simpset addsimps [sees_Cons] 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   922
                           setloop split_tac [expand_if]) 1); 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   923
by (asm_simp_tac (!simpset addsimps [parts_Un]) 1); 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   924
qed "in_parts_sees_Says";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   925
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   926
goal thy "!!evs. length evs < length evs' ==> newK  evs ~= newK (A',evs')";
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   927
by (fast_tac (!claset addSEs [less_irrefl]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   928
qed "length_less_newK_neq";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   929
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   930
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   931
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   932
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   933
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   934
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   935
YYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYYY;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   936
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   937
goalw thy [keysFor_def]
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   938
    "!!X. Crypt X K: H ==> invKey K : keysFor H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   939
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   940
qed "keysFor_I";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   941
AddSIs [keysFor_I];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   942
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   943
goalw thy [keysFor_def]
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   944
    "!!K. K : keysFor H ==> EX X K'. K=invKey K' & Crypt X K': H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   945
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   946
qed "keysFor_D";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   947
AddSDs [keysFor_D];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   948
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   949