src/HOL/Auth/NS_Public_Bad.ML
changeset 3121 cbb6c0c1c58a
parent 2637 e9b203f854ae
child 3440 22db7a9cbb52
--- a/src/HOL/Auth/NS_Public_Bad.ML	Wed May 07 12:50:26 1997 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.ML	Wed May 07 13:01:43 1997 +0200
@@ -16,15 +16,12 @@
 proof_timing:=true;
 HOL_quantifiers := false;
 
-val op addss = op unsafe_addss;
-
 AddIffs [Spy_in_lost];
 
 (*Replacing the variable by a constant improves search speed by 50%!*)
 val Says_imp_sees_Spy' = 
     read_instantiate_sg (sign_of thy) [("lost","lost")] Says_imp_sees_Spy;
 
-
 (*A "possibility property": there are traces that reach the end*)
 goal thy 
  "!!A B. A ~= B ==> EX NB. EX evs: ns_public.               \
@@ -46,17 +43,6 @@
 AddSEs   [not_Says_to_self RSN (2, rev_notE)];
 
 
-(*For proving the easier theorems about X ~: parts (sees lost Spy evs) *)
-fun parts_induct_tac i = SELECT_GOAL
-    (DETERM (etac ns_public.induct 1 THEN 
-             (*Fake message*)
-             TRY (best_tac (!claset addDs [impOfSubs analz_subset_parts,
-                                           impOfSubs Fake_parts_insert]
-                                    addss (!simpset)) 2)) THEN
-     (*Base case*)
-     fast_tac (!claset addss (!simpset)) 1 THEN
-     ALLGOALS Asm_simp_tac) i;
-
 (** Theorems of the form X ~: parts (sees lost Spy evs) imply that NOBODY
     sends messages containing X! **)
 
@@ -64,8 +50,9 @@
 goal thy 
  "!!evs. evs : ns_public \
 \        ==> (Key (priK A) : parts (sees lost Spy evs)) = (A : lost)";
-by (parts_induct_tac 1);
-by (Auto_tac());
+by (etac ns_public.induct 1);
+by (prove_simple_subgoals_tac 1);
+by (Fake_parts_insert_tac 1);
 qed "Spy_see_priK";
 Addsimps [Spy_see_priK];
 
@@ -78,7 +65,7 @@
 
 goal thy  "!!A. [| Key (priK A) : parts (sees lost Spy evs);       \
 \                  evs : ns_public |] ==> A:lost";
-by (fast_tac (!claset addDs [Spy_see_priK]) 1);
+by (blast_tac (!claset addDs [Spy_see_priK]) 1);
 qed "Spy_see_priK_D";
 
 bind_thm ("Spy_analz_priK_D", analz_subset_parts RS subsetD RS Spy_see_priK_D);
@@ -86,7 +73,7 @@
 
 
 fun analz_induct_tac i = 
-    etac ns_public.induct i     THEN
+    etac ns_public.induct i   THEN
     ALLGOALS (asm_simp_tac 
               (!simpset addsimps [not_parts_not_analz]
                         setloop split_tac [expand_if]));
@@ -105,16 +92,15 @@
 by (etac rev_mp 1);
 by (analz_induct_tac 1);
 (*NS3*)
-by (fast_tac (!claset addSEs partsEs) 4);
+by (blast_tac (!claset addSEs partsEs) 4);
 (*NS2*)
-by (fast_tac (!claset addSEs partsEs) 3);
+by (blast_tac (!claset addSEs partsEs) 3);
 (*Fake*)
-by (deepen_tac (!claset addSIs [analz_insertI]
+by (blast_tac (!claset addSIs [analz_insertI]
                         addDs [impOfSubs analz_subset_parts,
-			       impOfSubs Fake_parts_insert]
-			addss (!simpset)) 0 2);
+			       impOfSubs Fake_parts_insert]) 2);
 (*Base*)
-by (fast_tac (!claset addss (!simpset)) 1);
+by (Blast_tac 1);
 qed "no_nonce_NS1_NS2";
 
 
@@ -129,17 +115,16 @@
 (*NS1*)
 by (simp_tac (!simpset addsimps [all_conj_distrib]) 3);
 by (expand_case_tac "NA = ?y" 3 THEN
-    REPEAT (fast_tac (!claset addSEs partsEs) 3));
+    REPEAT (blast_tac (!claset addSEs partsEs) 3));
 (*Base*)
-by (fast_tac (!claset addss (!simpset)) 1);
+by (Blast_tac 1);
 (*Fake*)
 by (simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_sees]) 1);
 by (step_tac (!claset addSIs [analz_insertI]) 1);
 by (ex_strip_tac 1);
-by (best_tac (!claset delrules [conjI]
+by (blast_tac (!claset delrules [conjI]
                       addSDs [impOfSubs Fake_parts_insert]
-                      addDs  [impOfSubs analz_subset_parts]
-                      addss (!simpset)) 1);
+                      addDs  [impOfSubs analz_subset_parts]) 1);
 val lemma = result();
 
 goal thy 
@@ -160,15 +145,15 @@
 by (etac rev_mp 1);
 by (analz_induct_tac 1);
 (*NS3*)
-by (fast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
-                      addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 4);
+by (blast_tac (!claset addDs  [Says_imp_sees_Spy' RS parts.Inj]
+                       addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 4);
 (*NS2*)
-by (deepen_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
-                        addSEs [MPair_parts]
-			addDs  [parts.Body, unique_NA]) 0 3);
+by (blast_tac (!claset addSEs [MPair_parts]
+		       addDs  [Says_imp_sees_Spy' RS parts.Inj,
+			       parts.Body, unique_NA]) 3);
 (*NS1*)
-by (fast_tac (!claset addSEs sees_Spy_partsEs
-                      addIs  [impOfSubs analz_subset_parts]) 2);
+by (blast_tac (!claset addSEs sees_Spy_partsEs
+                       addIs  [impOfSubs analz_subset_parts]) 2);
 (*Fake*)
 by (spy_analz_tac 1);
 qed "Spy_not_see_NA";
@@ -187,20 +172,15 @@
 by (etac ns_public.induct 1);
 by (ALLGOALS Asm_simp_tac);
 (*NS1*)
-by (fast_tac (!claset addSEs sees_Spy_partsEs) 2);
+by (blast_tac (!claset addSEs sees_Spy_partsEs) 2);
 (*Fake*)
-by (REPEAT_FIRST (resolve_tac [impI, conjI]));
-by (fast_tac (!claset addss (!simpset)) 1);
-by (forward_tac [Spy_not_see_NA] 1 THEN REPEAT (assume_tac 1));
-by (best_tac (!claset addSIs [disjI2]
-                      addSDs [impOfSubs Fake_parts_insert]
-                      addDs  [impOfSubs analz_subset_parts]
-                      addss (!simpset)) 1);
-(*NS2*)
+by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert]
+                       addDs  [Spy_not_see_NA, 
+			       impOfSubs analz_subset_parts]) 1);
+(*NS2; not clear why blast_tac needs to be preceeded by Step_tac*)
 by (Step_tac 1);
-by (forward_tac [Spy_not_see_NA] 1 THEN REPEAT (assume_tac 1));
-by (deepen_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
-                        addDs  [unique_NA]) 1 1);
+by (blast_tac (!claset addDs [Says_imp_sees_Spy' RS parts.Inj,
+			      Spy_not_see_NA, unique_NA]) 1);
 qed "A_trusts_NS2";
 
 (*If the encrypted message appears then it originated with Alice in NS1*)
@@ -213,14 +193,12 @@
 by (etac rev_mp 1);
 by (analz_induct_tac 1);
 (*Fake*)
-by (best_tac (!claset addSIs [disjI2]
-                      addSDs [impOfSubs Fake_parts_insert]
-                      addIs  [analz_insertI]
-                      addDs  [impOfSubs analz_subset_parts]
-                      addss (!simpset)) 2);
+by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert]
+                       addIs  [analz_insertI]
+                       addDs  [impOfSubs analz_subset_parts]) 2);
 (*Base*)
-by (fast_tac (!claset addss (!simpset)) 1);
-qed_spec_mp "B_trusts_NS1";
+by (Blast_tac 1);
+qed "B_trusts_NS1";
 
 
 
@@ -238,17 +216,16 @@
 (*NS2*)
 by (simp_tac (!simpset addsimps [all_conj_distrib]) 3);
 by (expand_case_tac "NB = ?y" 3 THEN
-    REPEAT (fast_tac (!claset addSEs partsEs) 3));
+    REPEAT (blast_tac (!claset addSEs partsEs) 3));
 (*Base*)
-by (fast_tac (!claset addss (!simpset)) 1);
+by (Blast_tac 1);
 (*Fake*)
 by (simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_sees]) 1);
 by (step_tac (!claset addSIs [analz_insertI]) 1);
 by (ex_strip_tac 1);
-by (best_tac (!claset delrules [conjI]
+by (blast_tac (!claset delrules [conjI]
                       addSDs [impOfSubs Fake_parts_insert]
-                      addDs  [impOfSubs analz_subset_parts] 
-                      addss (!simpset)) 1);
+                      addDs  [impOfSubs analz_subset_parts]) 1);
 val lemma = result();
 
 goal thy 
@@ -271,22 +248,21 @@
 by (etac rev_mp 1);
 by (analz_induct_tac 1);
 (*NS1*)
-by (fast_tac (!claset addSEs sees_Spy_partsEs) 2);
+by (blast_tac (!claset addSEs sees_Spy_partsEs) 2);
 (*Fake*)
 by (spy_analz_tac 1);
 (*NS2 and NS3*)
 by (ALLGOALS (asm_simp_tac (!simpset addsimps [all_conj_distrib])));
 by (step_tac (!claset delrules [allI]) 1);
-by (Fast_tac 5);
-by (fast_tac (!claset addSIs [impOfSubs analz_subset_parts]) 1);
+by (Blast_tac 5);
+(*NS3*)
+by (blast_tac (!claset addDs [Says_imp_sees_Spy' RS parts.Inj,
+			      Spy_not_see_NB, unique_NB]) 4);
 (*NS2*)
-by (fast_tac (!claset addSEs sees_Spy_partsEs) 2);
-by (fast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
-                      addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 1);
-(*NS3*)
-by (forw_inst_tac [("A'","A")] (Says_imp_sees_Spy' RS parts.Inj RS unique_NB) 1
-    THEN REPEAT (eresolve_tac [asm_rl, Says_imp_sees_Spy' RS parts.Inj] 1));
-by (Fast_tac 1);
+by (blast_tac (!claset addSEs sees_Spy_partsEs) 3);
+by (blast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
+                       addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 2);
+by (blast_tac (!claset addSIs [impOfSubs analz_subset_parts]) 1);
 qed "Spy_not_see_NB";
 
 
@@ -305,22 +281,15 @@
 by (analz_induct_tac 1);
 by (ALLGOALS (asm_simp_tac (!simpset addsimps [ex_disj_distrib])));
 (*NS1*)
-by (fast_tac (!claset addSEs sees_Spy_partsEs) 2);
+by (blast_tac (!claset addSEs sees_Spy_partsEs) 2);
 (*Fake*)
-by (REPEAT_FIRST (resolve_tac [impI, conjI]));
-by (fast_tac (!claset addss (!simpset)) 1);
-by (rtac (ccontr RS disjI2) 1);
-by (forward_tac [Spy_not_see_NB] 1 THEN (REPEAT_FIRST assume_tac)
-    THEN Fast_tac 1);
-by (best_tac (!claset addSDs [impOfSubs Fake_parts_insert]
-                      addDs  [impOfSubs analz_subset_parts] 
-                      addss (!simpset)) 1);
-(*NS3*)
+by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert]
+                       addDs  [Spy_not_see_NB, 
+			       impOfSubs analz_subset_parts]) 1);
+(*NS3; not clear why blast_tac needs to be preceeded by Step_tac*)
 by (Step_tac 1);
-by (forward_tac [Spy_not_see_NB] 1 THEN (REPEAT_FIRST assume_tac)
-    THEN Fast_tac 1);
-by (best_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
-                      addDs  [unique_NB]) 1);
+by (blast_tac (!claset addDs [Says_imp_sees_Spy' RS parts.Inj,
+			      Spy_not_see_NB, unique_NB]) 1);
 qed "B_trusts_NS3";
 
 
@@ -331,18 +300,18 @@
 \     --> Nonce NB ~: analz (sees lost Spy evs)";
 by (analz_induct_tac 1);
 (*NS1*)
-by (fast_tac (!claset addSEs partsEs
-                      addSDs [Says_imp_sees_Spy' RS parts.Inj]) 2);
+by (blast_tac (!claset addSEs partsEs
+                       addSDs [Says_imp_sees_Spy' RS parts.Inj]) 2);
 (*Fake*)
 by (spy_analz_tac 1);
 (*NS2 and NS3*)
 by (Step_tac 1);
-by (fast_tac (!claset addSIs [impOfSubs analz_subset_parts, usedI]) 1);
+by (blast_tac (!claset addSIs [impOfSubs analz_subset_parts, usedI]) 1);
 (*NS2*)
-by (fast_tac (!claset addSEs partsEs
-                      addSDs [Says_imp_sees_Spy' RS parts.Inj]) 2);
-by (fast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
-                      addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 1);
+by (blast_tac (!claset addSEs partsEs
+                       addSDs [Says_imp_sees_Spy' RS parts.Inj]) 2);
+by (blast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj]
+                       addEs  [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 1);
 (*NS3*)
 by (forw_inst_tac [("A'","A")] (Says_imp_sees_Spy' RS parts.Inj RS unique_NB) 1
     THEN REPEAT (eresolve_tac [asm_rl, Says_imp_sees_Spy' RS parts.Inj] 1));