src/HOL/Auth/OtwayRees_Bad.ML
changeset 3102 4d01cdc119d2
parent 2837 dee1c7f1f576
child 3121 cbb6c0c1c58a
--- a/src/HOL/Auth/OtwayRees_Bad.ML	Mon May 05 12:15:20 1997 +0200
+++ b/src/HOL/Auth/OtwayRees_Bad.ML	Mon May 05 12:15:53 1997 +0200
@@ -20,6 +20,7 @@
 proof_timing:=true;
 HOL_quantifiers := false;
 
+(*No need to declare Says_imp_sees_Spy' because "lost" is a CONSTANT*)
 
 (*Weak liveness: there are traces that reach the end*)
 goal thy 
@@ -48,17 +49,17 @@
 
 goal thy "!!evs. Says A' B {|N, Agent A, Agent B, X|} : set_of_list evs ==> \
 \                X : analz (sees lost Spy evs)";
-by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]) 1);
+by (blast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]) 1);
 qed "OR2_analz_sees_Spy";
 
 goal thy "!!evs. Says S' B {|N, X, Crypt (shrK B) X'|} : set_of_list evs ==> \
 \                X : analz (sees lost Spy evs)";
-by (fast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]) 1);
+by (blast_tac (!claset addSDs [Says_imp_sees_Spy RS analz.Inj]) 1);
 qed "OR4_analz_sees_Spy";
 
 goal thy "!!evs. Says Server B {|NA, X, Crypt K' {|NB,K|}|} : set_of_list evs \
 \                 ==> K : parts (sees lost Spy evs)";
-by (fast_tac (!claset addSEs sees_Spy_partsEs) 1);
+by (blast_tac (!claset addSEs sees_Spy_partsEs) 1);
 qed "Oops_parts_sees_Spy";
 
 (*OR2_analz... and OR4_analz... let us treat those cases using the same 
@@ -109,7 +110,7 @@
 
 goal thy  "!!A. [| Key (shrK A) : parts (sees lost Spy evs);       \
 \                  evs : otway |] ==> A:lost";
-by (fast_tac (!claset addDs [Spy_see_shrK]) 1);
+by (blast_tac (!claset addDs [Spy_see_shrK]) 1);
 qed "Spy_see_shrK_D";
 
 bind_thm ("Spy_analz_shrK_D", analz_subset_parts RS subsetD RS Spy_see_shrK_D);
@@ -185,7 +186,7 @@
 by (REPEAT_FIRST (rtac analz_image_freshK_lemma ));
 by (ALLGOALS (asm_simp_tac analz_image_freshK_ss));
 (*Base*)
-by (fast_tac (!claset addIs [image_eqI] addss (!simpset)) 1);
+by (blast_tac (!claset addIs [image_eqI] addss (!simpset)) 1);
 (*Fake, OR2, OR4*) 
 by (REPEAT (spy_analz_tac 1));
 qed_spec_mp "analz_image_freshK";
@@ -211,11 +212,11 @@
 by (Step_tac 1);
 (*Remaining cases: OR3 and OR4*)
 by (ex_strip_tac 2);
-by (Fast_tac 2);
+by (Blast_tac 2);
 by (expand_case_tac "K = ?y" 1);
 by (REPEAT (ares_tac [refl,exI,impI,conjI] 2));
 (*...we assume X is a recent message, and handle this case by contradiction*)
-by (fast_tac (!claset addSEs sees_Spy_partsEs
+by (blast_tac (!claset addSEs sees_Spy_partsEs
                       delrules [conjI]    (*prevent split-up into 4 subgoals*)
                       addss (!simpset addsimps [parts_insertI])) 1);
 val lemma = result();
@@ -245,15 +246,13 @@
                             addsimps [not_parts_not_analz, analz_insert_freshK]
                             setloop split_tac [expand_if])));
 (*OR3*)
-by (fast_tac (!claset delrules [impCE]
+by (blast_tac (!claset delrules [impCE]
                       addSEs sees_Spy_partsEs
-                      addIs [impOfSubs analz_subset_parts]
-                      addss (!simpset addsimps [parts_insert2])) 3);
+                      addIs [impOfSubs analz_subset_parts]) 3);
 (*OR4, OR2, Fake*) 
 by (REPEAT_FIRST spy_analz_tac);
 (*Oops*) (** LEVEL 5 **)
-by (fast_tac (!claset delrules [disjE]
-                      addDs [unique_session_keys] addss (!simpset)) 1);
+by (blast_tac (!claset addSDs [unique_session_keys]) 1);
 val lemma = result() RS mp RS mp RSN(2,rev_notE);
 
 goal thy 
@@ -264,7 +263,7 @@
 \           A ~: lost;  B ~: lost;  evs : otway |]                \
 \        ==> Key K ~: analz (sees lost Spy evs)";
 by (forward_tac [Says_Server_message_form] 1 THEN assume_tac 1);
-by (fast_tac (!claset addSEs [lemma]) 1);
+by (blast_tac (!claset addSEs [lemma]) 1);
 qed "Spy_not_see_encrypted_key";
 
 
@@ -282,7 +281,7 @@
 \                       Crypt (shrK A) {|NA, Agent A, Agent B|}|}  \
 \             : set_of_list evs";
 by (parts_induct_tac 1);
-by (Fast_tac 1);
+by (Blast_tac 1);
 qed_spec_mp "Crypt_imp_OR1";
 
 
@@ -303,14 +302,13 @@
 \                   : set_of_list evs)";
 by (parts_induct_tac 1);
 (*OR1: it cannot be a new Nonce, contradiction.*)
-by (fast_tac (!claset addSIs [parts_insertI]
-                      addSEs sees_Spy_partsEs
-                      addss (!simpset)) 1);
+by (blast_tac (!claset addSIs [parts_insertI]
+                      addSEs sees_Spy_partsEs) 1);
 (*OR4*)
 by (REPEAT (Safe_step_tac 2));
-by (REPEAT (best_tac (!claset addSDs [parts_cut]) 3));
-by (fast_tac (!claset addSIs [Crypt_imp_OR1]
-                      addEs  sees_Spy_partsEs) 2);
+by (REPEAT (blast_tac (!claset addSDs [parts_cut]) 3));
+by (blast_tac (!claset addSIs [Crypt_imp_OR1]
+                       addEs  sees_Spy_partsEs) 2);
 (*OR3*)  (** LEVEL 5 **)
 by (ALLGOALS (asm_simp_tac (!simpset addsimps [ex_disj_distrib])));
 by (step_tac (!claset delrules [disjCI, impCE]) 1);