src/HOL/Auth/Shared.ML
author paulson
Mon, 23 Sep 1996 18:19:38 +0200
changeset 2012 1b234f1fd9c7
parent 2000 adb88d42f1bd
child 2026 0df5a96bf77e
permissions -rw-r--r--
Removal of the Notes constructor
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" 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     9
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    10
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
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    13
Addsimps [parts_cut_eq];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    14
2000
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
    15
(*GOALS.ML??*)
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
    16
fun prlim n = (goals_limit:=n; pr());
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
    17
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    18
(*FUN.ML??  WE NEED A NOTION OF INVERSE IMAGE, OR GRAPH!!*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    19
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
    20
by (fast_tac (!claset addEs [equalityE]) 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    21
val subset_range_iff = result();
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    22
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    23
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    24
open Shared;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    25
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    26
(*Holds because Friend is injective: thus cannot prove for all f*)
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    27
goal thy "(Friend x : Friend``A) = (x:A)";
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    28
by (Auto_tac());
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    29
qed "Friend_image_eq";
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    30
Addsimps [Friend_image_eq];
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    31
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    32
Addsimps [Un_insert_left, Un_insert_right];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    33
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    34
(*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
    35
  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
    36
  will not!*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    37
Addsimps [o_def];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    38
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    39
(*** Basic properties of shrK and newK ***)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    40
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    41
(* invKey (shrK A) = shrK A *)
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    42
bind_thm ("invKey_shrK", rewrite_rule [isSymKey_def] isSym_shrK);
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    43
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    44
(* invKey (newK evs) = newK evs *)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    45
bind_thm ("invKey_newK", rewrite_rule [isSymKey_def] isSym_newK);
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    46
Addsimps [invKey_shrK, invKey_newK];
1934
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
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
    49
(*Injectiveness and freshness of new keys and nonces*)
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
    50
AddIffs [inj_shrK RS inj_eq, inj_newN RS inj_eq, inj_newK RS inj_eq,
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
    51
	 fresh_newN, fresh_newK];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    52
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    53
(** Rewrites should not refer to  initState(Friend i) 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    54
    -- not in normal form! **)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    55
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    56
goal thy "newK evs ~= shrK B";
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    57
by (subgoal_tac "newK evs = shrK B --> \
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    58
\                Key (newK evs) : parts (initState B)" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    59
by (Fast_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    60
by (agent.induct_tac "B" 1);
1967
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1964
diff changeset
    61
by (auto_tac (!claset addIs [range_eqI], !simpset addsimps [bad_def]));
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    62
qed "newK_neq_shrK";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    63
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    64
Addsimps [newK_neq_shrK, newK_neq_shrK RS not_sym];
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
(*Good for talking about Server's initial state*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    67
goal thy "!!H. H <= Key``E ==> parts H = H";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    68
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    69
be parts.induct 1;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    70
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    71
qed "parts_image_subset";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    72
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    73
bind_thm ("parts_image_Key", subset_refl RS parts_image_subset);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    74
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    75
goal thy "!!H. H <= Key``E ==> analz H = H";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    76
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    77
be analz.induct 1;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    78
by (ALLGOALS (fast_tac (!claset addss (!simpset))));
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    79
qed "analz_image_subset";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    80
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    81
bind_thm ("analz_image_Key", subset_refl RS analz_image_subset);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    82
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    83
Addsimps [parts_image_Key, analz_image_Key];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    84
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    85
goalw thy [keysFor_def] "keysFor (parts (initState C)) = {}";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    86
by (agent.induct_tac "C" 1);
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
    87
by (auto_tac (!claset addIs [range_eqI], !simpset));
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    88
qed "keysFor_parts_initState";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    89
Addsimps [keysFor_parts_initState];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    90
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    91
goalw thy [keysFor_def] "keysFor (Key``E) = {}";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    92
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    93
qed "keysFor_image_Key";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    94
Addsimps [keysFor_image_Key];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    95
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    96
goal thy "shrK A ~: newK``E";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    97
by (agent.induct_tac "A" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    98
by (Auto_tac ());
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    99
qed "shrK_notin_image_newK";
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
   100
Addsimps [shrK_notin_image_newK];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   101
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
   102
(*Agents see their own shared keys!*)
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
   103
goal thy "Key (shrK A) : sees A evs";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   104
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   105
by (agent.induct_tac "A" 1);
1967
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1964
diff changeset
   106
by (auto_tac (!claset, !simpset addsimps [bad_def]));
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
   107
qed "sees_own_shrK";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   108
1967
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1964
diff changeset
   109
(*Enemy sees shared keys of bad agents!*)
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1964
diff changeset
   110
goal thy "!!i. A: bad ==> Key (shrK A) : sees Enemy evs";
1964
d551e68b7a36 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
   111
by (list.induct_tac "evs" 1);
1967
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1964
diff changeset
   112
by (auto_tac (!claset, !simpset addsimps [bad_def]));
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1964
diff changeset
   113
qed "Enemy_sees_bad";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   114
1967
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1964
diff changeset
   115
AddSIs [sees_own_shrK, Enemy_sees_bad];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   116
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   117
goal thy "Enemy: bad";
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   118
by (simp_tac (!simpset addsimps [bad_def]) 1);
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   119
qed "Enemy_in_bad";
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   120
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   121
AddIffs [Enemy_in_bad];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   122
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   123
(** Specialized rewrite rules for (sees A (Says...#evs)) **)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   124
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   125
goal thy "sees A (Says A B X # evs) = insert X (sees A evs)";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   126
by (Simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   127
qed "sees_own";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   128
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   129
goal thy "!!A. Server ~= A ==> \
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   130
\              sees Server (Says A B X # evs) = sees Server evs";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   131
by (Asm_simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   132
qed "sees_Server";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   133
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   134
goal thy "!!A. Friend i ~= A ==> \
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   135
\              sees (Friend i) (Says A B X # evs) = sees (Friend i) evs";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   136
by (Asm_simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   137
qed "sees_Friend";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   138
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   139
goal thy "sees Enemy (Says A B X # evs) = insert X (sees Enemy evs)";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   140
by (Simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   141
qed "sees_Enemy";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   142
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   143
goal thy "sees A (Says A' B X # evs) <= insert X (sees A evs)";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   144
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   145
by (Fast_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   146
qed "sees_Says_subset_insert";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   147
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   148
goal thy "sees A evs <= sees A (Says A' B X # evs)";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   149
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   150
by (Fast_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   151
qed "sees_subset_sees_Says";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   152
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   153
(*Pushing Unions into parts; one of the A's equals B, and thus sees Y*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   154
goal thy "(UN A. parts (sees A (Says B C Y # evs))) = \
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   155
\         parts {Y} Un (UN A. parts (sees A evs))";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   156
by (Step_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   157
be rev_mp 1;	(*for some reason, split_tac does not work on assumptions*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   158
val ss = (!simpset addsimps [parts_Un, sees_Cons] 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   159
	           setloop split_tac [expand_if]);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   160
by (ALLGOALS (fast_tac (!claset addss ss)));
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   161
qed "UN_parts_sees_Says";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   162
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   163
goal thy "Says A B X : set_of_list evs --> X : sees Enemy evs";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   164
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   165
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   166
qed_spec_mp "Says_imp_sees_Enemy";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   167
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   168
Addsimps [Says_imp_sees_Enemy];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   169
AddIs    [Says_imp_sees_Enemy];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   170
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
   171
goal thy "initState C <= Key `` range shrK";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   172
by (agent.induct_tac "C" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   173
by (Auto_tac ());
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   174
qed "initState_subset";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   175
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   176
goal thy "X : sees C evs --> \
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   177
\          (EX A B. Says A B X : set_of_list evs) | \
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
   178
\          (EX A. X = Key (shrK A))";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   179
by (list.induct_tac "evs" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   180
by (ALLGOALS Asm_simp_tac);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   181
by (fast_tac (!claset addDs [impOfSubs initState_subset]) 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   182
br conjI 1;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   183
by (Fast_tac 2);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   184
by (event.induct_tac "a" 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   185
by (ALLGOALS (asm_simp_tac (!simpset addsimps [mem_if])));
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   186
by (ALLGOALS Fast_tac);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   187
qed_spec_mp "seesD";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   188
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   189
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   190
Addsimps [UN_parts_sees_Says, sees_own, sees_Server, sees_Friend, sees_Enemy];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   191
Delsimps [sees_Cons];	(**** NOTE REMOVAL -- laws above are cleaner ****)
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
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   194
goal thy "!!K. newK evs = invKey K ==> newK evs = K";
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   195
br (invKey_eq RS iffD1) 1;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   196
by (Simp_tac 1);
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   197
val newK_invKey = result();
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   198
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   199
AddSDs [newK_invKey];
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   200
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   201
(** 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
   202
    be pulled out using the analz_insert rules **)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   203
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   204
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
   205
                      insert_commute;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   206
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   207
val pushKeys = map (insComm "Key ?K") 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   208
                  ["Agent ?C", "Nonce ?N", "MPair ?X ?Y", "Crypt ?X ?K'"];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   209
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   210
val pushCrypts = map (insComm "Crypt ?X ?K") 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   211
                    ["Agent ?C", "Nonce ?N", "MPair ?X' ?Y"];
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   212
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   213
(*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
   214
val pushes = pushKeys@pushCrypts;
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   215
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
   216
val pushKey_newK = insComm "Key (newK ?evs)"  "Key (shrK ?C)";
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   217
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
   218
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   219
(*No premature instantiation of variables.  For proving weak liveness.*)
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   220
fun safe_solver prems =
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   221
    match_tac (TrueI::refl::prems) ORELSE' eq_assume_tac
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   222
    ORELSE' etac FalseE;
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   223
2000
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
   224
(*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
   225
  Abstraction over i is ESSENTIAL: it delays the dereferencing of claset*)
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
   226
fun enemy_analz_tac i =
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   227
  SELECT_GOAL 
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   228
   (EVERY [  (*push in occurrences of X...*)
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   229
	   (REPEAT o CHANGED)
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   230
	     (res_inst_tac [("x1","X")] (insert_commute RS ssubst) 1),
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   231
	     (*...allowing further simplifications*)
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   232
	   simp_tac (!simpset setloop split_tac [expand_if]) 1,
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   233
	   REPEAT (resolve_tac [impI,notI] 1),
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   234
	   dtac (impOfSubs Fake_analz_insert) 1,
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   235
	   eresolve_tac [asm_rl, synth.Inj] 1,
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   236
	   Fast_tac 1,
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   237
	   Asm_full_simp_tac 1,
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   238
	   IF_UNSOLVED (deepen_tac (!claset addIs [impOfSubs analz_mono]) 0 1)
2000
adb88d42f1bd Abstraction of enemy_analz_tac over its argument
paulson
parents: 1993
diff changeset
   239
	   ]) i;
1993
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   240
77e7ef8e5c3b Addition of enemy_analz_tac and safe_solver
paulson
parents: 1967
diff changeset
   241
2012
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   242
(** Simplifying   parts (insert X (sees A evs))
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   243
                = parts {X} Un parts (sees A evs) -- since general case loops*)
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   244
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   245
val parts_insert_sees = 
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   246
    parts_insert |> read_instantiate_sg (sign_of thy) [("H", "sees A evs")]
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   247
                 |> standard;
1b234f1fd9c7 Removal of the Notes constructor
paulson
parents: 2000
diff changeset
   248