src/HOL/Auth/Message.ML
author paulson
Thu, 11 Jul 1996 15:30:22 +0200
changeset 1852 289ce6cb5c84
parent 1839 199243afac2b
child 1885 a18a6dc14f76
permissions -rw-r--r--
Added Msg 3; works up to Says_Server_imp_Key_newK
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
Datatypes of agents and messages;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     7
Inductive relations "parts", "analyze" and "synthesize"
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     8
*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     9
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    10
open Message;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    11
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    12
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    13
(**************** INSTALL CENTRALLY SOMEWHERE? ****************)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    14
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    15
(*Maybe swap the safe_tac and simp_tac lines?**)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    16
fun auto_tac (cs,ss) = 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    17
    TRY (safe_tac cs) THEN 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    18
    ALLGOALS (asm_full_simp_tac ss) THEN
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    19
    REPEAT (FIRSTGOAL (best_tac (cs addss ss)));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    20
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    21
fun Auto_tac() = auto_tac (!claset, !simpset);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    22
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    23
fun auto() = by (Auto_tac());
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    24
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
    25
fun impOfSubs th = th RSN (2, rev_subsetD);
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    26
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    27
(**************** INSTALL CENTRALLY SOMEWHERE? ****************)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    28
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    29
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    30
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    31
(** Inverse of keys **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    32
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    33
goal thy "!!K K'. (invKey K = invKey K') = (K=K')";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    34
by (Step_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    35
br box_equals 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    36
by (REPEAT (rtac invKey 2));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    37
by (Asm_simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    38
qed "invKey_eq";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    39
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    40
Addsimps [invKey, invKey_eq];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    41
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    42
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    43
(**** keysFor operator ****)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    44
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    45
goalw thy [keysFor_def] "keysFor {} = {}";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    46
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    47
qed "keysFor_empty";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    48
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    49
goalw thy [keysFor_def] "keysFor (H Un H') = keysFor H Un keysFor H'";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    50
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    51
qed "keysFor_Un";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    52
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    53
goalw thy [keysFor_def] "keysFor (UN i. H i) = (UN i. keysFor (H i))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    54
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    55
qed "keysFor_UN";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    56
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    57
(*Monotonicity*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    58
goalw thy [keysFor_def] "!!G H. G<=H ==> keysFor(G) <= keysFor(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    59
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    60
qed "keysFor_mono";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    61
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    62
goalw thy [keysFor_def] "keysFor (insert (Agent A) H) = keysFor H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    63
by (fast_tac (!claset addss (!simpset)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    64
qed "keysFor_insert_Agent";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    65
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    66
goalw thy [keysFor_def] "keysFor (insert (Nonce N) H) = keysFor H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    67
by (fast_tac (!claset addss (!simpset)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    68
qed "keysFor_insert_Nonce";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    69
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    70
goalw thy [keysFor_def] "keysFor (insert (Key K) H) = keysFor H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    71
by (fast_tac (!claset addss (!simpset)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    72
qed "keysFor_insert_Key";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    73
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    74
goalw thy [keysFor_def] "keysFor (insert {|X,Y|} H) = keysFor H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    75
by (fast_tac (!claset addss (!simpset)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    76
qed "keysFor_insert_MPair";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    77
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    78
goalw thy [keysFor_def]
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    79
    "keysFor (insert (Crypt X K) H) = insert (invKey K) (keysFor H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    80
by (Auto_tac());
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    81
by (fast_tac (!claset addIs [image_eqI]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    82
qed "keysFor_insert_Crypt";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    83
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    84
Addsimps [keysFor_empty, keysFor_Un, keysFor_UN, 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    85
	  keysFor_insert_Agent, keysFor_insert_Nonce,
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    86
	  keysFor_insert_Key, keysFor_insert_MPair,
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    87
	  keysFor_insert_Crypt];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    88
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    89
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    90
(**** Inductive relation "parts" ****)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    91
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    92
val major::prems = 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    93
goal thy "[| {|X,Y|} : parts H;       \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    94
\            [| X : parts H; Y : parts H |] ==> P  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    95
\         |] ==> P";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    96
by (cut_facts_tac [major] 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    97
brs prems 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    98
by (REPEAT (eresolve_tac [asm_rl, parts.Fst, parts.Snd] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    99
qed "MPair_parts";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   100
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   101
AddIs  [parts.Inj];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   102
AddSEs [MPair_parts];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   103
AddDs  [parts.Body];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   104
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   105
goal thy "H <= parts(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   106
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   107
qed "parts_increasing";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   108
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   109
(*Monotonicity*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   110
goalw thy parts.defs "!!G H. G<=H ==> parts(G) <= parts(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   111
by (rtac lfp_mono 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   112
by (REPEAT (ares_tac basic_monos 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   113
qed "parts_mono";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   114
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   115
goal thy "parts{} = {}";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   116
by (Step_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   117
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   118
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   119
qed "parts_empty";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   120
Addsimps [parts_empty];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   121
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   122
goal thy "!!X. X: parts{} ==> P";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   123
by (Asm_full_simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   124
qed "parts_emptyE";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   125
AddSEs [parts_emptyE];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   126
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   127
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   128
(** Unions **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   129
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   130
goal thy "parts(G) Un parts(H) <= parts(G Un H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   131
by (REPEAT (ares_tac [Un_least, parts_mono, Un_upper1, Un_upper2] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   132
val parts_Un_subset1 = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   133
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   134
goal thy "parts(G Un H) <= parts(G) Un parts(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   135
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   136
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   137
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   138
val parts_Un_subset2 = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   139
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   140
goal thy "parts(G Un H) = parts(G) Un parts(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   141
by (REPEAT (ares_tac [equalityI, parts_Un_subset1, parts_Un_subset2] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   142
qed "parts_Un";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   143
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   144
(*TWO inserts to avoid looping.  This rewrite is better than nothing...*)
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   145
goal thy "parts (insert X (insert Y H)) = parts {X} Un parts {Y} Un parts H";
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   146
by (stac (read_instantiate [("A","H")] insert_is_Un) 1);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   147
by (stac (read_instantiate [("A","{Y} Un H")] insert_is_Un) 1);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   148
by (simp_tac (HOL_ss addsimps [parts_Un, Un_assoc]) 1);
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   149
qed "parts_insert2";
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   150
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   151
goal thy "(UN x:A. parts(H x)) <= parts(UN x:A. H x)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   152
by (REPEAT (ares_tac [UN_least, parts_mono, UN_upper] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   153
val parts_UN_subset1 = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   154
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   155
goal thy "parts(UN x:A. H x) <= (UN x:A. parts(H x))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   156
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   157
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   158
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   159
val parts_UN_subset2 = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   160
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   161
goal thy "parts(UN x:A. H x) = (UN x:A. parts(H x))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   162
by (REPEAT (ares_tac [equalityI, parts_UN_subset1, parts_UN_subset2] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   163
qed "parts_UN";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   164
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   165
goal thy "parts(UN x. H x) = (UN x. parts(H x))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   166
by (simp_tac (!simpset addsimps [UNION1_def, parts_UN]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   167
qed "parts_UN1";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   168
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   169
(*Added to simplify arguments to parts, analyze and synthesize*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   170
Addsimps [parts_Un, parts_UN, parts_UN1];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   171
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   172
goal thy "insert X (parts H) <= parts(insert X H)";
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   173
by (fast_tac (!claset addEs [impOfSubs parts_mono]) 1);
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   174
qed "parts_insert_subset";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   175
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   176
(*Especially for reasoning about the Fake rule in traces*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   177
goal thy "!!Y. X: G ==> parts(insert X H) <= parts G Un parts H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   178
br ([parts_mono, parts_Un_subset2] MRS subset_trans) 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   179
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   180
qed "parts_insert_subset_Un";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   181
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   182
(** Idempotence and transitivity **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   183
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   184
goal thy "!!H. X: parts (parts H) ==> X: parts H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   185
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   186
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   187
qed "parts_partsE";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   188
AddSEs [parts_partsE];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   189
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   190
goal thy "parts (parts H) = parts H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   191
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   192
qed "parts_idem";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   193
Addsimps [parts_idem];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   194
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   195
goal thy "!!H. [| X: parts G;  G <= parts H |] ==> X: parts H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   196
by (dtac parts_mono 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   197
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   198
qed "parts_trans";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   199
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   200
(*Cut*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   201
goal thy "!!H. [| X: parts H;  Y: parts (insert X H) |] ==> Y: parts H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   202
be parts_trans 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   203
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   204
qed "parts_cut";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   205
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   206
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   207
(** Rewrite rules for pulling out atomic messages **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   208
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   209
goal thy "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   210
by (rtac (parts_insert_subset RSN (2, equalityI)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   211
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   212
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   213
(*Simplification breaks up equalities between messages;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   214
  how to make it work for fast_tac??*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   215
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   216
qed "parts_insert_Agent";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   217
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   218
goal thy "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   219
by (rtac (parts_insert_subset RSN (2, equalityI)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   220
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   221
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   222
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   223
qed "parts_insert_Nonce";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   224
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   225
goal thy "parts (insert (Key K) H) = insert (Key K) (parts H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   226
by (rtac (parts_insert_subset RSN (2, equalityI)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   227
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   228
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   229
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   230
qed "parts_insert_Key";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   231
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   232
goal thy "parts (insert (Crypt X K) H) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   233
\         insert (Crypt X K) (parts (insert X H))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   234
br equalityI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   235
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   236
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   237
by (Auto_tac());
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   238
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   239
by (ALLGOALS (best_tac (!claset addIs [parts.Body])));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   240
qed "parts_insert_Crypt";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   241
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   242
goal thy "parts (insert {|X,Y|} H) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   243
\         insert {|X,Y|} (parts (insert X (insert Y H)))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   244
br equalityI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   245
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   246
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   247
by (Auto_tac());
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   248
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   249
by (ALLGOALS (best_tac (!claset addIs [parts.Fst, parts.Snd])));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   250
qed "parts_insert_MPair";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   251
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   252
Addsimps [parts_insert_Agent, parts_insert_Nonce, 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   253
	  parts_insert_Key, parts_insert_Crypt, parts_insert_MPair];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   254
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   255
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   256
(**** Inductive relation "analyze" ****)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   257
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   258
val major::prems = 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   259
goal thy "[| {|X,Y|} : analyze H;       \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   260
\            [| X : analyze H; Y : analyze H |] ==> P  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   261
\         |] ==> P";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   262
by (cut_facts_tac [major] 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   263
brs prems 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   264
by (REPEAT (eresolve_tac [asm_rl, analyze.Fst, analyze.Snd] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   265
qed "MPair_analyze";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   266
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   267
AddIs  [analyze.Inj];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   268
AddSEs [MPair_analyze];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   269
AddDs  [analyze.Decrypt];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   270
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   271
goal thy "H <= analyze(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   272
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   273
qed "analyze_increasing";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   274
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   275
goal thy "analyze H <= parts H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   276
by (rtac subsetI 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   277
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   278
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   279
qed "analyze_subset_parts";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   280
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   281
bind_thm ("not_parts_not_analyze", analyze_subset_parts RS contra_subsetD);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   282
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   283
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   284
goal thy "parts (analyze H) = parts H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   285
br equalityI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   286
br (analyze_subset_parts RS parts_mono RS subset_trans) 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   287
by (Simp_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   288
by (fast_tac (!claset addDs [analyze_increasing RS parts_mono RS subsetD]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   289
qed "parts_analyze";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   290
Addsimps [parts_analyze];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   291
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   292
(*Monotonicity; Lemma 1 of Lowe*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   293
goalw thy analyze.defs "!!G H. G<=H ==> analyze(G) <= analyze(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   294
by (rtac lfp_mono 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   295
by (REPEAT (ares_tac basic_monos 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   296
qed "analyze_mono";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   297
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   298
(** General equational properties **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   299
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   300
goal thy "analyze{} = {}";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   301
by (Step_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   302
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   303
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   304
qed "analyze_empty";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   305
Addsimps [analyze_empty];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   306
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   307
(*Converse fails: we can analyze more from the union than from the 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   308
  separate parts, as a key in one might decrypt a message in the other*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   309
goal thy "analyze(G) Un analyze(H) <= analyze(G Un H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   310
by (REPEAT (ares_tac [Un_least, analyze_mono, Un_upper1, Un_upper2] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   311
qed "analyze_Un";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   312
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   313
goal thy "insert X (analyze H) <= analyze(insert X H)";
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   314
by (fast_tac (!claset addEs [impOfSubs analyze_mono]) 1);
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   315
qed "analyze_insert";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   316
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   317
(** Rewrite rules for pulling out atomic messages **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   318
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   319
goal thy "analyze (insert (Agent agt) H) = insert (Agent agt) (analyze H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   320
by (rtac (analyze_insert RSN (2, equalityI)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   321
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   322
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   323
(*Simplification breaks up equalities between messages;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   324
  how to make it work for fast_tac??*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   325
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   326
qed "analyze_insert_Agent";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   327
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   328
goal thy "analyze (insert (Nonce N) H) = insert (Nonce N) (analyze H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   329
by (rtac (analyze_insert RSN (2, equalityI)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   330
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   331
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   332
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   333
qed "analyze_insert_Nonce";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   334
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   335
(*Can only pull out Keys if they are not needed to decrypt the rest*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   336
goalw thy [keysFor_def]
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   337
    "!!K. K ~: keysFor (analyze H) ==>  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   338
\         analyze (insert (Key K) H) = insert (Key K) (analyze H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   339
by (rtac (analyze_insert RSN (2, equalityI)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   340
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   341
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   342
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   343
qed "analyze_insert_Key";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   344
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   345
goal thy "!!H. Key (invKey K) ~: analyze H ==>  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   346
\              analyze (insert (Crypt X K) H) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   347
\              insert (Crypt X K) (analyze H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   348
by (rtac (analyze_insert RSN (2, equalityI)) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   349
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   350
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   351
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   352
qed "analyze_insert_Crypt";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   353
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   354
goal thy "!!H. Key (invKey K) : analyze H ==>  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   355
\              analyze (insert (Crypt X K) H) <= \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   356
\              insert (Crypt X K) (analyze (insert X H))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   357
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   358
by (eres_inst_tac [("za","x")] analyze.induct 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   359
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   360
val lemma1 = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   361
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   362
goal thy "!!H. Key (invKey K) : analyze H ==>  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   363
\              insert (Crypt X K) (analyze (insert X H)) <= \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   364
\              analyze (insert (Crypt X K) H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   365
by (Auto_tac());
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   366
by (eres_inst_tac [("za","x")] analyze.induct 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   367
by (Auto_tac());
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   368
by (best_tac (!claset addIs [subset_insertI RS analyze_mono RS subsetD,
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   369
			     analyze.Decrypt]) 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   370
val lemma2 = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   371
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   372
goal thy "!!H. Key (invKey K) : analyze H ==>  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   373
\              analyze (insert (Crypt X K) H) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   374
\              insert (Crypt X K) (analyze (insert X H))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   375
by (REPEAT (ares_tac [equalityI, lemma1, lemma2] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   376
qed "analyze_insert_Decrypt";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   377
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   378
Addsimps [analyze_insert_Agent, analyze_insert_Nonce, 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   379
	  analyze_insert_Key, analyze_insert_Crypt,
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   380
	  analyze_insert_Decrypt];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   381
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   382
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   383
(*This rule supposes "for the sake of argument" that we have the key.*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   384
goal thy  "analyze (insert (Crypt X K) H) <=  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   385
\         insert (Crypt X K) (analyze (insert X H))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   386
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   387
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   388
by (Auto_tac());
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   389
qed "analyze_insert_Crypt_subset";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   390
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   391
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   392
(** Rewrite rules for pulling out atomic parts of messages **)
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 "analyze (insert X H) <= analyze (insert {|X,Y|} H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   395
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   396
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   397
by (ALLGOALS (best_tac (!claset addIs [analyze.Fst]))); 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   398
qed "analyze_insert_subset_MPair1";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   399
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   400
goal thy "analyze (insert Y H) <= analyze (insert {|X,Y|} H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   401
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   402
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   403
by (ALLGOALS (best_tac (!claset addIs [analyze.Snd]))); 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   404
qed "analyze_insert_subset_MPair2";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   405
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   406
goal thy "analyze (insert {|Agent agt,Y|} H) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   407
\         insert {|Agent agt,Y|} (insert (Agent agt) (analyze (insert Y H)))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   408
by (rtac equalityI 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   409
by (best_tac (!claset addIs [analyze.Fst,
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   410
			     impOfSubs analyze_insert_subset_MPair2]) 2); 
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   411
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   412
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   413
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   414
qed "analyze_insert_Agent_MPair";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   415
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   416
goal thy "analyze (insert {|Nonce N,Y|} H) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   417
\         insert {|Nonce N,Y|} (insert (Nonce N) (analyze (insert Y H)))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   418
by (rtac equalityI 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   419
by (best_tac (!claset addIs [analyze.Fst,
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   420
			     impOfSubs analyze_insert_subset_MPair2]) 2); 
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   421
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   422
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   423
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   424
qed "analyze_insert_Nonce_MPair";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   425
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   426
(*Can only pull out Keys if they are not needed to decrypt the rest*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   427
goalw thy [keysFor_def]
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   428
    "!!K. K ~: keysFor (analyze (insert Y H)) ==>  \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   429
\         analyze (insert {|Key K, Y|} H) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   430
\         insert {|Key K, Y|} (insert (Key K) (analyze (insert Y H)))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   431
by (rtac equalityI 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   432
by (best_tac (!claset addIs [analyze.Fst,
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   433
			     impOfSubs analyze_insert_subset_MPair2]) 2); 
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   434
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   435
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   436
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   437
qed "analyze_insert_Key_MPair";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   438
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   439
Addsimps [analyze_insert_Agent_MPair, analyze_insert_Nonce_MPair,
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   440
	  analyze_insert_Key_MPair];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   441
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   442
(** Idempotence and transitivity **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   443
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   444
goal thy "!!H. X: analyze (analyze H) ==> X: analyze H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   445
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   446
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   447
qed "analyze_analyzeE";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   448
AddSEs [analyze_analyzeE];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   449
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   450
goal thy "analyze (analyze H) = analyze H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   451
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   452
qed "analyze_idem";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   453
Addsimps [analyze_idem];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   454
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   455
goal thy "!!H. [| X: analyze G;  G <= analyze H |] ==> X: analyze H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   456
by (dtac analyze_mono 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   457
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   458
qed "analyze_trans";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   459
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   460
(*Cut; Lemma 2 of Lowe*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   461
goal thy "!!H. [| X: analyze H;  Y: analyze (insert X H) |] ==> Y: analyze H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   462
be analyze_trans 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   463
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   464
qed "analyze_cut";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   465
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   466
(*Cut can be proved easily by induction on
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   467
   "!!H. Y: analyze (insert X H) ==> X: analyze H --> Y: analyze H"
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   468
*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   469
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   470
(*If there are no pairs or encryptions then analyze does nothing*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   471
goal thy "!!H. [| ALL X Y. {|X,Y|} ~: H;  ALL X K. Crypt X K ~: H |] ==> \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   472
\         analyze H = H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   473
by (Step_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   474
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   475
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   476
qed "analyze_trivial";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   477
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   478
(*Helps to prove Fake cases*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   479
goal thy "!!X. X: analyze (UN i. analyze (H i)) ==> X: analyze (UN i. H i)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   480
be analyze.induct 1;
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   481
by (ALLGOALS (fast_tac (!claset addEs [impOfSubs analyze_mono])));
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   482
val lemma = result();
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   483
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   484
goal thy "analyze (UN i. analyze (H i)) = analyze (UN i. H i)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   485
by (fast_tac (!claset addIs [lemma]
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   486
		      addEs [impOfSubs analyze_mono]) 1);
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   487
qed "analyze_UN_analyze";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   488
Addsimps [analyze_UN_analyze];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   489
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   490
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   491
(**** Inductive relation "synthesize" ****)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   492
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   493
AddIs  synthesize.intrs;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   494
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   495
goal thy "H <= synthesize(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   496
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   497
qed "synthesize_increasing";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   498
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   499
(*Monotonicity*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   500
goalw thy synthesize.defs "!!G H. G<=H ==> synthesize(G) <= synthesize(H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   501
by (rtac lfp_mono 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   502
by (REPEAT (ares_tac basic_monos 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   503
qed "synthesize_mono";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   504
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   505
(** Unions **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   506
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   507
(*Converse fails: we can synthesize more from the union than from the 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   508
  separate parts, building a compound message using elements of each.*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   509
goal thy "synthesize(G) Un synthesize(H) <= synthesize(G Un H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   510
by (REPEAT (ares_tac [Un_least, synthesize_mono, Un_upper1, Un_upper2] 1));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   511
qed "synthesize_Un";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   512
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   513
(** Idempotence and transitivity **)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   514
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   515
goal thy "!!H. X: synthesize (synthesize H) ==> X: synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   516
be synthesize.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   517
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   518
qed "synthesize_synthesizeE";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   519
AddSEs [synthesize_synthesizeE];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   520
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   521
goal thy "synthesize (synthesize H) = synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   522
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   523
qed "synthesize_idem";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   524
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   525
goal thy "!!H. [| X: synthesize G;  G <= synthesize H |] ==> X: synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   526
by (dtac synthesize_mono 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   527
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   528
qed "synthesize_trans";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   529
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   530
(*Cut; Lemma 2 of Lowe*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   531
goal thy "!!H. [| X: synthesize H;  Y: synthesize (insert X H) \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   532
\              |] ==> Y: synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   533
be synthesize_trans 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   534
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   535
qed "synthesize_cut";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   536
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   537
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   538
(*Can only produce a nonce or key if it is already known,
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   539
  but can synthesize a pair or encryption from its components...*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   540
val mk_cases = synthesize.mk_cases msg.simps;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   541
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   542
val Nonce_synthesize = mk_cases "Nonce n : synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   543
val Key_synthesize   = mk_cases "Key K : synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   544
val MPair_synthesize = mk_cases "{|X,Y|} : synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   545
val Crypt_synthesize = mk_cases "Crypt X K : synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   546
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   547
AddSEs [Nonce_synthesize, Key_synthesize, MPair_synthesize, Crypt_synthesize];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   548
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   549
goal thy "(Nonce N : synthesize H) = (Nonce N : H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   550
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   551
qed "Nonce_synthesize_eq";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   552
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   553
goal thy "(Key K : synthesize H) = (Key K : H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   554
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   555
qed "Key_synthesize_eq";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   556
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   557
Addsimps [Nonce_synthesize_eq, Key_synthesize_eq];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   558
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   559
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   560
goalw thy [keysFor_def]
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   561
    "keysFor (synthesize H) = keysFor H Un invKey``{K. Key K : H}";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   562
by (Fast_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   563
qed "keysFor_synthesize";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   564
Addsimps [keysFor_synthesize];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   565
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   566
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   567
(*** Combinations of parts, analyze and synthesize ***)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   568
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   569
(*Not that useful, in view of the following one...*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   570
goal thy "parts (synthesize H) <= synthesize (parts H)";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   571
by (Step_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   572
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   573
be (parts_increasing RS synthesize_mono RS subsetD) 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   574
by (ALLGOALS Fast_tac);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   575
qed "parts_synthesize_subset";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   576
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   577
goal thy "parts (synthesize H) = parts H Un synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   578
br equalityI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   579
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   580
be parts.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   581
by (ALLGOALS
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   582
    (best_tac (!claset addIs ((synthesize_increasing RS parts_mono RS subsetD)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   583
			     ::parts.intrs))));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   584
qed "parts_synthesize";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   585
Addsimps [parts_synthesize];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   586
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   587
goal thy "analyze (synthesize H) = analyze H Un synthesize H";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   588
br equalityI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   589
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   590
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   591
by (best_tac
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   592
    (!claset addIs [synthesize_increasing RS analyze_mono RS subsetD]) 5);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   593
(*Strange that best_tac just can't hack this one...*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   594
by (ALLGOALS (deepen_tac (!claset addIs analyze.intrs) 0));
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   595
qed "analyze_synthesize";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   596
Addsimps [analyze_synthesize];
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   597
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   598
(*Hard to prove; still needed now that there's only one Enemy?*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   599
goal thy "analyze (UN i. synthesize (H i)) = \
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   600
\         analyze (UN i. H i) Un (UN i. synthesize (H i))";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   601
br equalityI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   602
br subsetI 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   603
be analyze.induct 1;
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   604
by (best_tac
1852
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   605
    (!claset addEs [impOfSubs synthesize_increasing,
289ce6cb5c84 Added Msg 3; works up to Says_Server_imp_Key_newK
paulson
parents: 1839
diff changeset
   606
		    impOfSubs analyze_mono]) 5);
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   607
by (Best_tac 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   608
by (deepen_tac (!claset addIs [analyze.Fst]) 0 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   609
by (deepen_tac (!claset addIs [analyze.Snd]) 0 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   610
by (deepen_tac (!claset addSEs [analyze.Decrypt]
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   611
			addIs  [analyze.Decrypt]) 0 1);
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   612
qed "analyze_UN1_synthesize";
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   613
Addsimps [analyze_UN1_synthesize];