src/HOL/Auth/Shared.ML
author paulson
Wed, 09 Apr 1997 12:32:04 +0200
changeset 2922 580647a879cf
parent 2891 d8f254ad1ab9
child 3121 cbb6c0c1c58a
permissions -rw-r--r--
Using Blast_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2320
41289907faed Moved much common material to Message.ML
paulson
parents: 2284
diff changeset
     1
(*  Title:      HOL/Auth/Shared
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     2
    ID:         $Id$
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     5
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     6
Theory of Shared Keys (common to all symmetric-key protocols)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     7
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     8
Server keys; initial states of agents; new nonces and keys; function "sees" 
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
     9
*)
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    10
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    11
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    12
open Shared;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    13
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    14
(*Holds because Friend is injective: thus cannot prove for all f*)
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    15
goal thy "(Friend x : Friend``A) = (x:A)";
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    16
by (Auto_tac());
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    17
qed "Friend_image_eq";
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    18
Addsimps [Friend_image_eq];
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    19
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    20
Addsimps [Un_insert_left, Un_insert_right];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    21
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    22
(*By default only o_apply is built-in.  But in the presence of eta-expansion
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    23
  this means that some terms displayed as (f o g) will be rewritten, and others
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    24
  will not!*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    25
Addsimps [o_def];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    26
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    27
(*** Basic properties of shrK and newK ***)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    28
2451
ce85a2aafc7a Extensive tidying and simplification, largely stemming from
paulson
parents: 2376
diff changeset
    29
(*Injectiveness and freshness of new keys and nonces*)
ce85a2aafc7a Extensive tidying and simplification, largely stemming from
paulson
parents: 2376
diff changeset
    30
AddIffs [inj_shrK RS inj_eq, inj_newN RS inj_eq, 
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    31
         inj_newK RS inj_eq, inj_nPair RS inj_eq];
2376
f5c61fd9b9b6 Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents: 2320
diff changeset
    32
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    33
(* invKey (shrK A) = shrK A *)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    34
bind_thm ("invKey_id", rewrite_rule [isSymKey_def] isSym_keys);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    35
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    36
Addsimps [invKey_id];
2376
f5c61fd9b9b6 Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents: 2320
diff changeset
    37
2451
ce85a2aafc7a Extensive tidying and simplification, largely stemming from
paulson
parents: 2376
diff changeset
    38
goal thy "!!K. newK i = invKey K ==> newK i = K";
2376
f5c61fd9b9b6 Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents: 2320
diff changeset
    39
by (rtac (invKey_eq RS iffD1) 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    40
by (Full_simp_tac 1);
2376
f5c61fd9b9b6 Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents: 2320
diff changeset
    41
val newK_invKey = result();
f5c61fd9b9b6 Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents: 2320
diff changeset
    42
f5c61fd9b9b6 Temporary additions (random) for the nested Otway-Rees protocol
paulson
parents: 2320
diff changeset
    43
AddSDs [newK_invKey, sym RS newK_invKey];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    44
2451
ce85a2aafc7a Extensive tidying and simplification, largely stemming from
paulson
parents: 2376
diff changeset
    45
Addsimps [newK_neq_shrK, newK_neq_shrK RS not_sym];
2320
41289907faed Moved much common material to Message.ML
paulson
parents: 2284
diff changeset
    46
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    47
(** Rewrites should not refer to  initState(Friend i) 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    48
    -- not in normal form! **)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    49
2451
ce85a2aafc7a Extensive tidying and simplification, largely stemming from
paulson
parents: 2376
diff changeset
    50
goal thy "Key (newK i) ~: parts (initState lost B)";
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    51
by (agent.induct_tac "B" 1);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    52
by (Auto_tac ());
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    53
qed "newK_notin_initState";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    54
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    55
AddIffs [newK_notin_initState];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    56
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    57
goalw thy [keysFor_def] "keysFor (parts (initState lost C)) = {}";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    58
by (agent.induct_tac "C" 1);
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
    59
by (auto_tac (!claset addIs [range_eqI], !simpset));
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    60
qed "keysFor_parts_initState";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    61
Addsimps [keysFor_parts_initState];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    62
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    63
goalw thy [keysFor_def] "keysFor (Key``E) = {}";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    64
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    65
qed "keysFor_image_Key";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    66
Addsimps [keysFor_image_Key];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    67
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    68
goal thy "shrK A ~: newK``E";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    69
by (agent.induct_tac "A" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    70
by (Auto_tac ());
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    71
qed "shrK_notin_image_newK";
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    72
Addsimps [shrK_notin_image_newK];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    73
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    74
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    75
(*** Function "sees" ***)
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    76
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    77
goal thy
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    78
    "!!evs. lost' <= lost ==> sees lost' A evs <= sees lost A evs";
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    79
by (list.induct_tac "evs" 1);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    80
by (agent.induct_tac "A" 1);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    81
by (event.induct_tac "a" 2);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    82
by (Auto_tac ());
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    83
qed "sees_mono";
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    84
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    85
(*Agents see their own shared keys!*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    86
goal thy "A ~= Spy --> Key (shrK A) : sees lost A evs";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    87
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    88
by (agent.induct_tac "A" 1);
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    89
by (Auto_tac ());
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    90
qed_spec_mp "sees_own_shrK";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    91
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    92
(*Spy sees shared keys of lost agents!*)
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
    93
goal thy "!!A. A: lost ==> Key (shrK A) : sees lost Spy evs";
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    94
by (list.induct_tac "evs" 1);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    95
by (Auto_tac());
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
    96
qed "Spy_sees_lost";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    97
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
    98
AddSIs [sees_own_shrK, Spy_sees_lost];
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    99
2124
9677fdf5fc23 New theorem Crypt_Spy_analz_lost; improvements to spy_analz_tac; ex_strip_tac
paulson
parents: 2109
diff changeset
   100
(*Added for Yahalom/lost_tac*)
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2264
diff changeset
   101
goal thy "!!A. [| Crypt (shrK A) X : analz (sees lost Spy evs);  A: lost |] \
2124
9677fdf5fc23 New theorem Crypt_Spy_analz_lost; improvements to spy_analz_tac; ex_strip_tac
paulson
parents: 2109
diff changeset
   102
\              ==> X : analz (sees lost Spy evs)";
9677fdf5fc23 New theorem Crypt_Spy_analz_lost; improvements to spy_analz_tac; ex_strip_tac
paulson
parents: 2109
diff changeset
   103
by (fast_tac (!claset addSDs [analz.Decrypt] addss (!simpset)) 1);
9677fdf5fc23 New theorem Crypt_Spy_analz_lost; improvements to spy_analz_tac; ex_strip_tac
paulson
parents: 2109
diff changeset
   104
qed "Crypt_Spy_analz_lost";
9677fdf5fc23 New theorem Crypt_Spy_analz_lost; improvements to spy_analz_tac; ex_strip_tac
paulson
parents: 2109
diff changeset
   105
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   106
(** Specialized rewrite rules for (sees lost A (Says...#evs)) **)
2026
0df5a96bf77e Last working version prior to introduction of "lost"
paulson
parents: 2012
diff changeset
   107
2078
b198b3d46fb4 New version of axiom sees1_Says:
paulson
parents: 2072
diff changeset
   108
goal thy "sees lost B (Says A B X # evs) = insert X (sees lost B evs)";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   109
by (Simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   110
qed "sees_own";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   111
2078
b198b3d46fb4 New version of axiom sees1_Says:
paulson
parents: 2072
diff changeset
   112
goal thy "!!A. Server ~= B ==> \
b198b3d46fb4 New version of axiom sees1_Says:
paulson
parents: 2072
diff changeset
   113
\          sees lost Server (Says A B X # evs) = sees lost Server evs";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   114
by (Asm_simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   115
qed "sees_Server";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   116
2078
b198b3d46fb4 New version of axiom sees1_Says:
paulson
parents: 2072
diff changeset
   117
goal thy "!!A. Friend i ~= B ==> \
b198b3d46fb4 New version of axiom sees1_Says:
paulson
parents: 2072
diff changeset
   118
\          sees lost (Friend i) (Says A B X # evs) = sees lost (Friend i) evs";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   119
by (Asm_simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   120
qed "sees_Friend";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   121
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   122
goal thy "sees lost Spy (Says A B X # evs) = insert X (sees lost Spy evs)";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   123
by (Simp_tac 1);
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   124
qed "sees_Spy";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   125
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   126
goal thy "sees lost A (Says A' B X # evs) <= insert X (sees lost A evs)";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   127
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   128
by (Blast_tac 1);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   129
qed "sees_Says_subset_insert";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   130
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   131
goal thy "sees lost A evs <= sees lost A (Says A' B X # evs)";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   132
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   133
by (Blast_tac 1);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   134
qed "sees_subset_sees_Says";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   135
2132
aeba09ebd8bc Tidied up a big mess in UN_parts_sees_Says
paulson
parents: 2124
diff changeset
   136
(*Pushing Unions into parts.  One of the agents A is B, and thus sees Y.
2160
ad4382e546fc Simplified new_keys_not_seen, etc.: replaced the
paulson
parents: 2132
diff changeset
   137
  Once used to prove new_keys_not_seen; now obsolete.*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   138
goal thy "(UN A. parts (sees lost A (Says B C Y # evs))) = \
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   139
\         parts {Y} Un (UN A. parts (sees lost A evs))";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   140
by (Step_tac 1);
2132
aeba09ebd8bc Tidied up a big mess in UN_parts_sees_Says
paulson
parents: 2124
diff changeset
   141
by (etac rev_mp 1);     (*split_tac does not work on assumptions*)
aeba09ebd8bc Tidied up a big mess in UN_parts_sees_Says
paulson
parents: 2124
diff changeset
   142
by (ALLGOALS
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2516
diff changeset
   143
    (fast_tac (!claset unsafe_addss (!simpset addsimps [parts_Un, sees_Cons] 
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   144
				            setloop split_tac [expand_if]))));
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   145
qed "UN_parts_sees_Says";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   146
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   147
goal thy "Says A B X : set_of_list evs --> X : sees lost Spy evs";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   148
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   149
by (Auto_tac ());
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   150
qed_spec_mp "Says_imp_sees_Spy";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   151
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   152
(*Use with addSEs to derive contradictions from old Says events containing
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   153
  items known to be fresh*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   154
val sees_Spy_partsEs = make_elim (Says_imp_sees_Spy RS parts.Inj):: partsEs;
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   155
2064
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   156
goal thy  
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2264
diff changeset
   157
 "!!evs. [| Says A B (Crypt (shrK C) X) : set_of_list evs;  C : lost |] \
2064
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   158
\        ==> X : analz (sees lost Spy evs)";
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   159
by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2516
diff changeset
   160
                      unsafe_addss (!simpset)) 1);
2064
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   161
qed "Says_Crypt_lost";
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   162
2072
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   163
goal thy  
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2264
diff changeset
   164
 "!!evs. [| Says A B (Crypt (shrK C) X) : set_of_list evs;        \
2072
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   165
\           X ~: analz (sees lost Spy evs) |]                     \
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   166
\        ==> C ~: lost";
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   167
by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2516
diff changeset
   168
                      unsafe_addss (!simpset)) 1);
2072
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   169
qed "Says_Crypt_not_lost";
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   170
2320
41289907faed Moved much common material to Message.ML
paulson
parents: 2284
diff changeset
   171
(*NEEDED??*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   172
goal thy "initState lost C <= Key `` range shrK";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   173
by (agent.induct_tac "C" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   174
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   175
qed "initState_subset";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   176
2320
41289907faed Moved much common material to Message.ML
paulson
parents: 2284
diff changeset
   177
(*NEEDED??*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   178
goal thy "X : sees lost C evs --> \
2078
b198b3d46fb4 New version of axiom sees1_Says:
paulson
parents: 2072
diff changeset
   179
\          (EX A B. Says A B X : set_of_list evs) | (EX A. X = Key (shrK A))";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   180
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   181
by (ALLGOALS Asm_simp_tac);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   182
by (blast_tac (!claset addDs [impOfSubs initState_subset]) 1);
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   183
by (rtac conjI 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   184
by (Blast_tac 2);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   185
by (event.induct_tac "a" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   186
by (ALLGOALS (asm_simp_tac (!simpset addsimps [mem_if])));
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   187
by (ALLGOALS Blast_tac);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   188
qed_spec_mp "seesD";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   189
2160
ad4382e546fc Simplified new_keys_not_seen, etc.: replaced the
paulson
parents: 2132
diff changeset
   190
Addsimps [sees_own, sees_Server, sees_Friend, sees_Spy];
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   191
Delsimps [sees_Cons];   (**** NOTE REMOVAL -- laws above are cleaner ****)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   192
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   193
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   194
(*** Fresh nonces ***)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   195
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   196
goal thy "Nonce N ~: parts (initState lost B)";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   197
by (agent.induct_tac "B" 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   198
by (Auto_tac ());
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   199
qed "Nonce_notin_initState";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   200
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   201
AddIffs [Nonce_notin_initState];
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   202
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   203
goalw thy [used_def] "!!X. X: parts (sees lost B evs) ==> X: used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   204
by (etac (impOfSubs parts_mono) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   205
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   206
qed "usedI";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   207
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   208
AddIs [usedI];
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   209
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   210
(** Fresh keys never clash with long-term shared keys **)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   211
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   212
goal thy "Key (shrK A) : used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   213
by (Best_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   214
qed "shrK_in_used";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   215
AddIffs [shrK_in_used];
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   216
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   217
(*Used in parts_Fake_tac and analz_Fake_tac to distinguish session keys
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   218
  from long-term shared keys*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   219
goal thy "!!K. Key K ~: used evs ==> K ~: range shrK";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   220
by (Best_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   221
qed "Key_not_used";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   222
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   223
(*A session key cannot clash with a long-term shared key*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   224
goal thy "!!K. K ~: range shrK ==> shrK B ~= K";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   225
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   226
qed "shrK_neq";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   227
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   228
Addsimps [Key_not_used, shrK_neq, shrK_neq RS not_sym];
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   229
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   230
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   231
goal thy "used (Says A B X # evs) = parts{X} Un used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   232
by (simp_tac (!simpset addsimps [used_def, UN_parts_sees_Says]) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   233
qed "used_Says";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   234
Addsimps [used_Says];
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   235
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   236
goal thy "used [] <= used l";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   237
by (list.induct_tac "l" 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   238
by (event.induct_tac "a" 2);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   239
by (ALLGOALS Asm_simp_tac);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   240
by (Best_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   241
qed "used_nil_subset";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   242
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   243
goal thy "used l <= used (l@l')";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   244
by (list.induct_tac "l" 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   245
by (simp_tac (!simpset addsimps [used_nil_subset]) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   246
by (event.induct_tac "a" 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   247
by (Asm_simp_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   248
by (Best_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   249
qed "used_subset_append";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   250
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   251
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   252
(*** Supply fresh nonces for possibility theorems. ***)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   253
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   254
goalw thy [used_def] "EX N. ALL n. N<=n --> Nonce n ~: used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   255
by (list.induct_tac "evs" 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   256
by (res_inst_tac [("x","0")] exI 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   257
by (Step_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   258
by (Full_simp_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   259
(*Inductive step*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   260
by (event.induct_tac "a" 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   261
by (full_simp_tac (!simpset addsimps [UN_parts_sees_Says]) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   262
by (msg.induct_tac "msg" 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   263
by (ALLGOALS (asm_simp_tac (!simpset addsimps [exI, parts_insert2])));
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   264
by (Step_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   265
(*MPair case*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   266
by (res_inst_tac [("x","Na+Nb")] exI 2);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   267
by (blast_tac (!claset addSEs [add_leE]) 2);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   268
(*Nonce case*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   269
by (res_inst_tac [("x","N + Suc nat")] exI 1);
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2516
diff changeset
   270
by (fast_tac (!claset addSEs [add_leE] addaltern trans_tac) 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   271
val lemma = result();
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   272
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   273
goal thy "EX N. Nonce N ~: used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   274
by (rtac (lemma RS exE) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   275
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   276
qed "Nonce_supply1";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   277
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   278
goal thy "EX N N'. Nonce N ~: used evs & Nonce N' ~: used evs' & N ~= N'";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   279
by (cut_inst_tac [("evs","evs")] lemma 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   280
by (cut_inst_tac [("evs","evs'")] lemma 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   281
by (Step_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   282
by (res_inst_tac [("x","N")] exI 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   283
by (res_inst_tac [("x","Suc (N+Na)")] exI 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   284
by (asm_simp_tac (!simpset addsimps [less_not_refl2 RS not_sym, 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   285
				     le_add2, le_add1, 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   286
				     le_eq_less_Suc RS sym]) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   287
qed "Nonce_supply2";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   288
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   289
goal thy "EX N N' N''. Nonce N ~: used evs & Nonce N' ~: used evs' & \
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   290
\                   Nonce N'' ~: used evs'' & N ~= N' & N' ~= N'' & N ~= N''";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   291
by (cut_inst_tac [("evs","evs")] lemma 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   292
by (cut_inst_tac [("evs","evs'")] lemma 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   293
by (cut_inst_tac [("evs","evs''")] lemma 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   294
by (Step_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   295
by (res_inst_tac [("x","N")] exI 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   296
by (res_inst_tac [("x","Suc (N+Na)")] exI 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   297
by (res_inst_tac [("x","Suc (Suc (N+Na+Nb))")] exI 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   298
by (asm_simp_tac (!simpset addsimps [less_not_refl2 RS not_sym, 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   299
				     le_add2, le_add1, 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   300
				     le_eq_less_Suc RS sym]) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   301
by (rtac (less_trans RS less_not_refl2 RS not_sym) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   302
by (stac (le_eq_less_Suc RS sym) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   303
by (asm_simp_tac (!simpset addsimps [le_eq_less_Suc RS sym]) 2);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   304
by (REPEAT (rtac le_add1 1));
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   305
qed "Nonce_supply3";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   306
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   307
goal thy "Nonce (@ N. Nonce N ~: used evs) ~: used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   308
by (rtac (lemma RS exE) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   309
by (rtac selectI 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   310
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   311
qed "Nonce_supply";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   312
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   313
(*** Supply fresh keys for possibility theorems. ***)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   314
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   315
goal thy "EX K. Key K ~: used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   316
by (rtac (Fin.emptyI RS Key_supply_ax RS exE) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   317
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   318
qed "Key_supply1";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   319
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   320
val Fin_UNIV_insertI = UNIV_I RS Fin.insertI;
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   321
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   322
goal thy "EX K K'. Key K ~: used evs & Key K' ~: used evs' & K ~= K'";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   323
by (cut_inst_tac [("evs","evs")] (Fin.emptyI RS Key_supply_ax) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   324
by (etac exE 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   325
by (cut_inst_tac [("evs","evs'")] 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   326
    (Fin.emptyI RS Fin_UNIV_insertI RS Key_supply_ax) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   327
by (Auto_tac());
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   328
qed "Key_supply2";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   329
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   330
goal thy "EX K K' K''. Key K ~: used evs & Key K' ~: used evs' & \
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   331
\                      Key K'' ~: used evs'' & K ~= K' & K' ~= K'' & K ~= K''";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   332
by (cut_inst_tac [("evs","evs")] (Fin.emptyI RS Key_supply_ax) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   333
by (etac exE 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   334
by (cut_inst_tac [("evs","evs'")] 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   335
    (Fin.emptyI RS Fin_UNIV_insertI RS Key_supply_ax) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   336
by (etac exE 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   337
by (cut_inst_tac [("evs","evs''")] 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   338
    (Fin.emptyI RS Fin_UNIV_insertI RS Fin_UNIV_insertI RS Key_supply_ax) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   339
by (Step_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   340
by (Full_simp_tac 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   341
by (fast_tac (!claset addSEs [allE]) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   342
qed "Key_supply3";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   343
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   344
goal thy "Key (@ K. Key K ~: used evs) ~: used evs";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   345
by (rtac (Fin.emptyI RS Key_supply_ax RS exE) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   346
by (rtac selectI 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   347
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   348
qed "Key_supply";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   349
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   350
(*** Tactics for possibility theorems ***)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   351
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   352
val possibility_tac =
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   353
    REPEAT (*omit used_Says so that Nonces, Keys start from different traces!*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   354
    (ALLGOALS (simp_tac 
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2516
diff changeset
   355
               (!simpset delsimps [used_Says] setSolver safe_solver))
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   356
     THEN
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   357
     REPEAT_FIRST (eq_assume_tac ORELSE' 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   358
                   resolve_tac [refl, conjI, Nonce_supply, Key_supply]));
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   359
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   360
(*For harder protocols (such as Recur) where we have to set up some
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   361
  nonces and keys initially*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   362
val basic_possibility_tac =
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   363
    REPEAT 
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2516
diff changeset
   364
    (ALLGOALS (asm_simp_tac (!simpset setSolver safe_solver))
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   365
     THEN
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   366
     REPEAT_FIRST (resolve_tac [refl, conjI]));
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   367
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   368
2049
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   369
(** Power of the Spy **)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   370
2049
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   371
(*The Spy can see more than anybody else, except for their initial state*)
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   372
goal thy "sees lost A evs <= initState lost A Un sees lost Spy evs";
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   373
by (list.induct_tac "evs" 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   374
by (event.induct_tac "a" 2);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   375
by (ALLGOALS (fast_tac (!claset addDs [sees_Says_subset_insert RS subsetD] 
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   376
                                addss (!simpset))));
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   377
qed "sees_agent_subset_sees_Spy";
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   378
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   379
(*The Spy can see more than anybody else who's lost their key!*)
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   380
goal thy "A: lost --> A ~= Server --> sees lost A evs <= sees lost Spy evs";
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   381
by (list.induct_tac "evs" 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   382
by (event.induct_tac "a" 2);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   383
by (agent.induct_tac "A" 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   384
by (auto_tac (!claset addDs [sees_Says_subset_insert RS subsetD], (!simpset)));
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   385
qed_spec_mp "sees_lost_agent_subset_sees_Spy";
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   386
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   387
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   388
(** Simplifying   parts (insert X (sees lost A evs))
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   389
      = parts {X} Un parts (sees lost A evs) -- since general case loops*)
2012
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   390
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   391
val parts_insert_sees = 
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   392
    parts_insert |> read_instantiate_sg (sign_of thy)
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   393
                                        [("H", "sees lost A evs")]
2012
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   394
                 |> standard;
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   395
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   396
2049
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   397
(*** Specialized rewriting for analz_insert_Key_newK ***)
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   398
2320
41289907faed Moved much common material to Message.ML
paulson
parents: 2284
diff changeset
   399
(*Push newK applications in, allowing other keys to be pulled out*)
41289907faed Moved much common material to Message.ML
paulson
parents: 2284
diff changeset
   400
val pushKey_newK = insComm thy "Key (newK ?evs)"  "Key (shrK ?C)";
41289907faed Moved much common material to Message.ML
paulson
parents: 2284
diff changeset
   401
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   402
goal thy "!!A. A <= Compl (range shrK) ==> shrK x ~: A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   403
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   404
qed "subset_Compl_range";
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   405
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   406
goal thy "insert (Key K) H = Key `` {K} Un H";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   407
by (Blast_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   408
qed "insert_Key_singleton";
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   409
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   410
goal thy "insert (Key K) (Key``KK Un C) = Key `` (insert K KK) Un C";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   411
by (Blast_tac 1);
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   412
qed "insert_Key_image";
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   413
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   414
val analz_image_freshK_ss = 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   415
     !simpset delsimps [image_insert, image_Un]
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   416
              addsimps ([image_insert RS sym, image_Un RS sym,
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   417
                         Key_not_used, 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   418
                         insert_Key_singleton, subset_Compl_range,
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   419
                         insert_Key_image, Un_assoc RS sym]
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   420
                        @disj_comms)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   421
              setloop split_tac [expand_if];
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   422
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   423
(*Lemma for the trivial direction of the if-and-only-if*)
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   424
goal thy  
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   425
 "!!evs. (Key K : analz (Key``nE Un H)) --> (K : nE | Key K : analz H)  ==> \
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   426
\        (Key K : analz (Key``nE Un H)) = (K : nE | Key K : analz H)";
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
   427
by (blast_tac (!claset addIs [impOfSubs analz_mono]) 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
   428
qed "analz_image_freshK_lemma";