Tidied up the unicity proofs
authorpaulson
Tue, 07 Jan 1997 10:18:20 +0100
changeset 2480 f9be937df511
parent 2479 57109c1a653d
child 2481 ee461c8bc9c3
Tidied up the unicity proofs
src/HOL/Auth/NS_Public.ML
src/HOL/Auth/NS_Public_Bad.ML
--- a/src/HOL/Auth/NS_Public.ML	Tue Jan 07 10:17:07 1997 +0100
+++ b/src/HOL/Auth/NS_Public.ML	Tue Jan 07 10:18:20 1997 +0100
@@ -84,9 +84,8 @@
 
 (*** Future nonces can't be seen or used! ***)
 
-goal thy "!!evs. evs : ns_public ==>          \
-\                length evs <= i --> \
-\                Nonce (newN i) ~: parts (sees lost Spy evs)";
+goal thy "!!i. evs : ns_public ==>        \
+\             length evs <= i --> Nonce (newN i) ~: parts (sees lost Spy evs)";
 by (parts_induct_tac 1);
 by (REPEAT_FIRST (fast_tac (!claset 
                               addSEs partsEs
@@ -115,7 +114,7 @@
  "!!evs. evs : ns_public                       \
 \ ==> Nonce NA ~: analz (sees lost Spy evs) -->           \
 \     Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs) --> \
-\     Crypt (pubK C) {|X, Nonce NA, Agent D|} ~: parts (sees lost Spy evs)";
+\     Crypt (pubK C) {|NA', Nonce NA, Agent D|} ~: parts (sees lost Spy evs)";
 by (analz_induct_tac 1);
 (*NS3*)
 by (fast_tac (!claset addSEs partsEs
@@ -133,7 +132,7 @@
 bind_thm ("no_nonce_NS1_NS2", result() RSN (2, rev_mp) RSN (2, rev_mp));
 
 
-(*Uniqueness for NS1: nonce NA identifies agents A and B*)
+(*Unicity for NS1: nonce NA identifies agents A and B*)
 goal thy 
  "!!evs. evs : ns_public                                                    \
 \ ==> Nonce NA ~: analz (sees lost Spy evs) -->                             \
@@ -162,12 +161,7 @@
 \           Nonce NA ~: analz (sees lost Spy evs);                            \
 \           evs : ns_public |]                                                \
 \        ==> A=A' & B=B'";
-by (dtac lemma 1);
-by (mp_tac 1);
-by (REPEAT (etac exE 1));
-(*Duplicate the assumption*)
-by (forw_inst_tac [("psi", "ALL C.?P(C)")] asm_rl 1);
-by (fast_tac (!claset addSDs [spec]) 1);
+by (prove_unique_tac lemma 1);
 qed "unique_NA";
 
 
@@ -240,7 +234,8 @@
  "!!evs. evs : ns_public                   \
 \    ==> Nonce NA ~: analz (sees lost Spy evs) --> \
 \        Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs) --> \
-\        Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs"; by (analz_induct_tac 1);
+\        Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set_of_list evs";
+by (analz_induct_tac 1);
 (*Fake*)
 by (best_tac (!claset addSIs [disjI2]
 		      addSDs [impOfSubs Fake_parts_insert]
@@ -255,7 +250,7 @@
 
 (**** Authenticity properties obtained from NS2 ****)
 
-(*Uniqueness for NS2: nonce NB identifies nonce NA and agents A, B 
+(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B 
   [unicity of B makes Lowe's fix work]
   [proof closely follows that for unique_NA] *)
 goal thy 
--- a/src/HOL/Auth/NS_Public_Bad.ML	Tue Jan 07 10:17:07 1997 +0100
+++ b/src/HOL/Auth/NS_Public_Bad.ML	Tue Jan 07 10:18:20 1997 +0100
@@ -26,9 +26,8 @@
 
 (*A "possibility property": there are traces that reach the end*)
 goal thy 
- "!!A B. [| A ~= B |]   \
-\        ==> EX NB. EX evs: ns_public.               \
-\               Says A B (Crypt (pubK B) (Nonce NB)) : set_of_list evs";
+ "!!A B. A ~= B ==> EX NB. EX evs: ns_public.               \
+\                     Says A B (Crypt (pubK B) (Nonce NB)) : set_of_list evs";
 by (REPEAT (resolve_tac [exI,bexI] 1));
 by (rtac (ns_public.Nil RS ns_public.NS1 RS ns_public.NS2 RS ns_public.NS3) 2);
 by (ALLGOALS (simp_tac (!simpset setsolver safe_solver)));
@@ -109,7 +108,7 @@
     etac ns_public.induct i	THEN
     ALLGOALS (asm_simp_tac 
 	      (!simpset addsimps [not_parts_not_analz]
-               setloop split_tac [expand_if]));
+                        setloop split_tac [expand_if]));
 
 
 (**** Authenticity properties obtained from NS2 ****)
@@ -138,7 +137,7 @@
 bind_thm ("no_nonce_NS1_NS2", result() RSN (2, rev_mp) RSN (2, rev_mp));
 
 
-(*Uniqueness for NS1: nonce NA identifies agents A and B*)
+(*Unicity for NS1: nonce NA identifies agents A and B*)
 goal thy 
  "!!evs. evs : ns_public                                                    \
 \ ==> Nonce NA ~: analz (sees lost Spy evs) -->                             \
@@ -257,7 +256,7 @@
 
 (**** Authenticity properties obtained from NS2 ****)
 
-(*Uniqueness for NS2: nonce NB identifies agent A and nonce NA
+(*Unicity for NS2: nonce NB identifies agent A and nonce NA
   [proof closely follows that for unique_NA] *)
 goal thy 
  "!!evs. evs : ns_public                                                    \