src/HOL/Auth/Shared.ML
author paulson
Tue, 08 Oct 1996 10:28:02 +0200
changeset 2072 6ac12b9478d5
parent 2064 5a5e508e2a2b
child 2078 b198b3d46fb4
permissions -rw-r--r--
Put in the theorem Says_Crypt_not_lost
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Message
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
2072
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
    12
Addsimps [de_Morgan_conj, de_Morgan_disj, not_all, not_ex];
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
    13
(**Addsimps [all_conj_distrib, ex_disj_distrib];**)
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
    14
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
    15
2064
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    16
val prems = goal HOL.thy "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)";
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    17
by (excluded_middle_tac "P" 1);
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    18
by (asm_simp_tac (!simpset addsimps prems) 1);
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    19
by (asm_simp_tac (!simpset addsimps prems) 1);
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    20
val expand_case = result();
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    21
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    22
fun expand_case_tac P i =
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    23
    res_inst_tac [("P",P)] expand_case i THEN
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    24
    Simp_tac (i+1) THEN 
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    25
    Simp_tac i;
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    26
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
    27
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    28
2000
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
    29
(*GOALS.ML??*)
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
    30
fun prlim n = (goals_limit:=n; pr());
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
    31
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    32
(*FUN.ML??  WE NEED A NOTION OF INVERSE IMAGE, OR GRAPH!!*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    33
goal Set.thy "!!f. B <= range f = (B = f`` {x. f x: B})";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    34
by (fast_tac (!claset addEs [equalityE]) 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    35
val subset_range_iff = result();
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    36
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    37
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    38
open Shared;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    39
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    40
(*Holds because Friend is injective: thus cannot prove for all f*)
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    41
goal thy "(Friend x : Friend``A) = (x:A)";
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    42
by (Auto_tac());
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    43
qed "Friend_image_eq";
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    44
Addsimps [Friend_image_eq];
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    45
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    46
Addsimps [Un_insert_left, Un_insert_right];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    47
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    48
(*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
    49
  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
    50
  will not!*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    51
Addsimps [o_def];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    52
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    53
(*** Basic properties of shrK and newK ***)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    54
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    55
(* invKey (shrK A) = shrK A *)
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    56
bind_thm ("invKey_shrK", rewrite_rule [isSymKey_def] isSym_shrK);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    57
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    58
(* invKey (newK evs) = newK evs *)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    59
bind_thm ("invKey_newK", rewrite_rule [isSymKey_def] isSym_newK);
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    60
Addsimps [invKey_shrK, invKey_newK];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    61
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    62
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
    63
(*Injectiveness and freshness of new keys and nonces*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    64
AddIffs [inj_shrK RS inj_eq, inj_newN RS inj_eq, inj_newK RS inj_eq];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    65
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    66
(** Rewrites should not refer to  initState(Friend i) 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    67
    -- not in normal form! **)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    68
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    69
Addsimps [newK_neq_shrK, newK_neq_shrK RS not_sym];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    70
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    71
goal thy "Key (newK evs) ~: parts (initState lost B)";
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    72
by (agent.induct_tac "B" 1);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    73
by (Auto_tac ());
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    74
qed "newK_notin_initState";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    75
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    76
goal thy "Nonce (newN evs) ~: parts (initState lost B)";
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    77
by (agent.induct_tac "B" 1);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    78
by (Auto_tac ());
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    79
qed "newN_notin_initState";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    80
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    81
AddIffs [newK_notin_initState, newN_notin_initState];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    82
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
    83
goalw thy [keysFor_def] "keysFor (parts (initState lost C)) = {}";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    84
by (agent.induct_tac "C" 1);
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
    85
by (auto_tac (!claset addIs [range_eqI], !simpset));
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    86
qed "keysFor_parts_initState";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    87
Addsimps [keysFor_parts_initState];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    88
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    89
goalw thy [keysFor_def] "keysFor (Key``E) = {}";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    90
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    91
qed "keysFor_image_Key";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    92
Addsimps [keysFor_image_Key];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    93
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    94
goal thy "shrK A ~: newK``E";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    95
by (agent.induct_tac "A" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    96
by (Auto_tac ());
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    97
qed "shrK_notin_image_newK";
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    98
Addsimps [shrK_notin_image_newK];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    99
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   100
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   101
(*** Function "sees" ***)
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   102
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   103
goal thy
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   104
    "!!evs. lost' <= lost ==> sees lost' A evs <= sees lost A evs";
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   105
by (list.induct_tac "evs" 1);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   106
by (agent.induct_tac "A" 1);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   107
by (event.induct_tac "a" 2);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   108
by (Auto_tac ());
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   109
qed "sees_mono";
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   110
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
   111
(*Agents see their own shared keys!*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   112
goal thy "A ~= Spy --> Key (shrK A) : sees lost A evs";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   113
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   114
by (agent.induct_tac "A" 1);
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   115
by (Auto_tac ());
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   116
qed_spec_mp "sees_own_shrK";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   117
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   118
(*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
   119
goal thy "!!A. A: lost ==> Key (shrK A) : sees lost Spy evs";
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   120
by (list.induct_tac "evs" 1);
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   121
by (Auto_tac());
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   122
qed "Spy_sees_lost";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   123
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   124
AddSIs [sees_own_shrK, Spy_sees_lost];
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   125
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   126
(** Specialized rewrite rules for (sees lost A (Says...#evs)) **)
2026
0df5a96bf77e Last working version prior to introduction of "lost"
paulson
parents: 2012
diff changeset
   127
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   128
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
   129
by (Simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   130
qed "sees_own";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   131
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   132
goal thy "!!A. Server ~= A ==> \
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   133
\              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
   134
by (Asm_simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   135
qed "sees_Server";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   136
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   137
goal thy "!!A. Friend i ~= A ==> \
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   138
\              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
   139
by (Asm_simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   140
qed "sees_Friend";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   141
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   142
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
   143
by (Simp_tac 1);
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   144
qed "sees_Spy";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   145
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   146
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
   147
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   148
by (Fast_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   149
qed "sees_Says_subset_insert";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   150
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   151
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
   152
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   153
by (Fast_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   154
qed "sees_subset_sees_Says";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   155
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   156
(*Pushing Unions into parts; one of the A's equals B, and thus sees lost Y*)
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   157
goal thy "(UN A. parts (sees lost A (Says B C Y # evs))) = \
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   158
\         parts {Y} Un (UN A. parts (sees lost A evs))";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   159
by (Step_tac 1);
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   160
by (etac rev_mp 1);     (*for some reason, split_tac does not work on assumptions*)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   161
val ss = (!simpset addsimps [parts_Un, sees_Cons] 
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   162
                   setloop split_tac [expand_if]);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   163
by (ALLGOALS (fast_tac (!claset addss ss)));
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   164
qed "UN_parts_sees_Says";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   165
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   166
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
   167
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   168
by (Auto_tac ());
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   169
qed_spec_mp "Says_imp_sees_Spy";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   170
2064
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   171
goal thy  
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   172
 "!!evs. [| Says A B (Crypt X (shrK C)) : set_of_list evs;        \
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   173
\           C   : lost |]                                         \
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   174
\        ==> X : analz (sees lost Spy evs)";
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   175
by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   176
                      addss (!simpset)) 1);
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   177
qed "Says_Crypt_lost";
5a5e508e2a2b Simple tidying
paulson
parents: 2049
diff changeset
   178
2072
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   179
goal thy  
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   180
 "!!evs. [| Says A B (Crypt X (shrK C)) : set_of_list evs;        \
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   181
\           X ~: analz (sees lost Spy evs) |]                     \
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   182
\        ==> C ~: lost";
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   183
by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   184
                      addss (!simpset)) 1);
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   185
qed "Says_Crypt_not_lost";
6ac12b9478d5 Put in the theorem Says_Crypt_not_lost
paulson
parents: 2064
diff changeset
   186
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   187
goal thy "initState lost C <= Key `` range shrK";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   188
by (agent.induct_tac "C" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   189
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   190
qed "initState_subset";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   191
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   192
goal thy "X : sees lost C evs --> \
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   193
\          (EX A B. Says A B X : set_of_list evs) | \
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
   194
\          (EX A. X = Key (shrK A))";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   195
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   196
by (ALLGOALS Asm_simp_tac);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   197
by (fast_tac (!claset addDs [impOfSubs initState_subset]) 1);
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   198
by (rtac conjI 1);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   199
by (Fast_tac 2);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   200
by (event.induct_tac "a" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   201
by (ALLGOALS (asm_simp_tac (!simpset addsimps [mem_if])));
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   202
by (ALLGOALS Fast_tac);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   203
qed_spec_mp "seesD";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   204
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   205
Addsimps [UN_parts_sees_Says, sees_own, sees_Server, sees_Friend, sees_Spy];
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   206
Delsimps [sees_Cons];   (**** NOTE REMOVAL -- laws above are cleaner ****)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   207
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   208
2049
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   209
(** Power of the Spy **)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   210
2049
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   211
(*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
   212
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
   213
by (list.induct_tac "evs" 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   214
by (event.induct_tac "a" 2);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   215
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
   216
                                addss (!simpset))));
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   217
qed "sees_agent_subset_sees_Spy";
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   218
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   219
(*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
   220
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
   221
by (list.induct_tac "evs" 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   222
by (event.induct_tac "a" 2);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   223
by (agent.induct_tac "A" 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   224
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
   225
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
   226
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   227
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   228
(** Rewrites to push in Key and Crypt messages, so that other messages can
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   229
    be pulled out using the analz_insert rules **)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   230
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   231
fun insComm x y = read_instantiate_sg (sign_of thy) [("x",x), ("y",y)] 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   232
                      insert_commute;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   233
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   234
val pushKeys = map (insComm "Key ?K") 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   235
                  ["Agent ?C", "Nonce ?N", "MPair ?X ?Y", "Crypt ?X ?K'"];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   236
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   237
val pushCrypts = map (insComm "Crypt ?X ?K") 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   238
                    ["Agent ?C", "Nonce ?N", "MPair ?X' ?Y"];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   239
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   240
(*Cannot be added with Addsimps -- we don't always want to re-order messages*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   241
val pushes = pushKeys@pushCrypts;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   242
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
   243
val pushKey_newK = insComm "Key (newK ?evs)"  "Key (shrK ?C)";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   244
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   245
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   246
(*No premature instantiation of variables.  For proving weak liveness.*)
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   247
fun safe_solver prems =
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   248
    match_tac (TrueI::refl::prems) ORELSE' eq_assume_tac
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   249
    ORELSE' etac FalseE;
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   250
2000
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
   251
(*Analysis of Fake cases and of messages that forward unknown parts
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
   252
  Abstraction over i is ESSENTIAL: it delays the dereferencing of claset*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   253
fun spy_analz_tac i =
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   254
  SELECT_GOAL 
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   255
   (EVERY [  (*push in occurrences of X...*)
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   256
           (REPEAT o CHANGED)
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   257
             (res_inst_tac [("x1","X")] (insert_commute RS ssubst) 1),
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   258
             (*...allowing further simplifications*)
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   259
           simp_tac (!simpset setloop split_tac [expand_if]) 1,
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   260
           REPEAT (resolve_tac [impI,notI] 1),
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   261
           dtac (impOfSubs Fake_analz_insert) 1,
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   262
           eresolve_tac [asm_rl, synth.Inj] 1,
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   263
           Fast_tac 1,
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   264
           Asm_full_simp_tac 1,
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   265
           IF_UNSOLVED (deepen_tac (!claset addIs [impOfSubs analz_mono]) 0 1)
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   266
           ]) i;
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   267
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   268
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2026
diff changeset
   269
(** 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
   270
      = parts {X} Un parts (sees lost A evs) -- since general case loops*)
2012
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   271
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   272
val parts_insert_sees = 
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   273
    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
   274
                                        [("H", "sees lost A evs")]
2012
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   275
                 |> standard;
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   276
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   277
2049
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   278
(*** Specialized rewriting for analz_insert_Key_newK ***)
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   279
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   280
goal thy "!!K. newK evs = invKey K ==> newK evs = K";
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   281
by (rtac (invKey_eq RS iffD1) 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   282
by (Simp_tac 1);
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   283
val newK_invKey = result();
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   284
d3b93e1765bc Moved sees_lost_agent_subset_sees_Spy to common file
paulson
parents: 2045
diff changeset
   285
AddSDs [newK_invKey];
2045
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   286
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   287
Delsimps [image_insert];
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   288
Addsimps [image_insert RS sym];
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   289
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   290
Delsimps [image_Un];
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   291
Addsimps [image_Un RS sym];
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   292
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   293
goal thy "insert (Key (newK x)) H = Key `` (newK``{x}) Un H";
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   294
by (Fast_tac 1);
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   295
qed "insert_Key_singleton";
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   296
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   297
goal thy "insert (Key (f x)) (Key``(f``E) Un C) = \
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   298
\         Key `` (f `` (insert x E)) Un C";
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   299
by (Fast_tac 1);
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   300
qed "insert_Key_image";
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   301
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   302
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   303
(*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
   304
goal thy  
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   305
 "!!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
   306
\        (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
   307
by (fast_tac (!claset addSEs [impOfSubs analz_mono]) 1);
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   308
qed "analz_image_newK_lemma";
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   309
ae1030e66745 Removed some dead wood. Transferred lemmas used to prove analz_image_newK
paulson
parents: 2032
diff changeset
   310