--- a/src/HOL/Auth/Event.ML Tue Aug 20 17:46:24 1996 +0200
+++ b/src/HOL/Auth/Event.ML Tue Aug 20 18:53:17 1996 +0200
@@ -384,14 +384,22 @@
be traces.induct 1;
bd NS3_msg_in_parts_sees_Enemy 5;
by (ALLGOALS Asm_simp_tac);
-by (ALLGOALS
- (best_tac
+(*NS1 and NS2*)
+map (by o fast_tac (!claset addDs [Suc_leD] addss (!simpset))) [3,2];
+(*Fake and NS3*)
+map (by o best_tac
(!claset addSDs [newK_invKey]
addDs [impOfSubs (analz_subset_parts RS keysFor_mono),
impOfSubs (parts_insert_subset_Un RS keysFor_mono),
Suc_leD]
addEs [new_keys_not_seen RS not_parts_not_analz RSN (2,rev_notE)]
- addss (!simpset))));
+ addss (!simpset)))
+ [2,1];
+(*NS4*)
+by (step_tac (!claset addSDs [Suc_leD, newK_invKey]) 1);
+by (deepen_tac (!claset addIs [impOfSubs keysFor_mono]) 0 2);
+by (fast_tac (!claset addDs [Says_imp_old_keys]
+ addss (!simpset addsimps [le_def])) 1);
val lemma = result();
goal thy
@@ -526,6 +534,7 @@
by (auto_tac (!claset addIs [Says_imp_old_keys], !simpset));
qed "Says_S_message_form";
+(*Currently unused. Needed only to reason about the VERY LAST message.*)
goal thy
"!!evs. [| evs = Says S A (Crypt {|Nonce NA, Agent B, Key K, X|} \
\ (serverKey A)) # evs'; \
@@ -540,7 +549,6 @@
-
(****
The following is to prove theorems of the form
@@ -568,6 +576,7 @@
addDs [impOfSubs analz_subset_parts,
impOfSubs parts_insert_subset_Un]
addss (!simpset)) 1);
+by (fast_tac (!claset addss (!simpset)) 1);
result();
--- a/src/HOL/Auth/Event.thy Tue Aug 20 17:46:24 1996 +0200
+++ b/src/HOL/Auth/Event.thy Tue Aug 20 18:53:17 1996 +0200
@@ -85,24 +85,27 @@
(*NS3 DOESN'T ALLOW INTERLEAVING -- that is, it only responds to the
MOST RECENT message.*)
+(*Needham-Schroeder Shared-Key protocol (from BAN paper, page 247)*)
consts traces :: "event list set"
inductive traces
intrs
+ (*Initial trace is empty*)
Nil "[]: traces"
- (*The enemy MAY say anything he CAN say. We do not expect him to
- invent new nonces here, but he can also use NS1.*)
- Fake "[| evs: traces; B ~= Enemy;
- X: synth(analz(sees Enemy evs))
+ (*The enemy MAY say anything he CAN say. We do not expect him to
+ invent new nonces here, but he can also use NS1. Common to
+ all similar protocols.*)
+ Fake "[| evs: traces; B ~= Enemy; X: synth (analz (sees Enemy evs))
|] ==> (Says Enemy B X) # evs : traces"
+ (*Alice initiates a protocol run*)
NS1 "[| evs: traces; A ~= Server
|] ==> (Says A Server {|Agent A, Agent B, Nonce (newN evs)|})
# evs : traces"
- (*We can't trust the sender field, hence the A' in it.
- This allows the Server to respond more than once to A's
- request...*)
+ (*Server's response to Alice's message.
+ !! It may respond more than once to A's request !!
+ We can't trust the sender field, hence the A' in it.*)
NS2 "[| evs: traces; A ~= B; A ~= Server;
(Says A' Server {|Agent A, Agent B, Nonce NA|}) : set_of_list evs
|] ==> (Says Server A
@@ -110,9 +113,9 @@
(Crypt {|Key (newK evs), Agent A|} (serverKey B))|}
(serverKey A))) # evs : traces"
- (*We can't assume S=Server. Agent A "remembers" her nonce.
- May assume WLOG that she is NOT the Enemy: the Fake rule
- covers this case. Can inductively show A ~= Server*)
+ (*We can't assume S=Server. Agent A "remembers" her nonce.
+ May assume WLOG that she is NOT the Enemy: the Fake rule
+ covers this case. Can inductively show A ~= Server*)
NS3 "[| evs: traces; A ~= B;
(Says S A (Crypt {|Nonce NA, Agent B, Key K, X|} (serverKey A)))
: set_of_list evs;
@@ -120,12 +123,8 @@
(Says A Server {|Agent A, Agent B, Nonce NA|}) : set_of_list evs
|] ==> (Says A B X) # evs : traces"
-(*Initial version of NS2 had
-
- {|Agent A, Agent B, Key (newK evs), Nonce NA|}
-
- for the encrypted part; the version above is LESS explicit, hence
- HARDER to prove. Also it is precisely as in the BAN paper.
-*)
-
+ NS4 "[| evs: traces; A ~= B;
+ (Says A' B (Crypt {|Key K, Agent A|} (serverKey B)))
+ : set_of_list evs
+ |] ==> (Says B A (Crypt (Nonce (newN evs)) K)) # evs : traces"
end