src/ZF/Trancl.ML
changeset 2929 4eefc6c22d41
parent 2493 bdeb5024353a
child 3016 15763781afb0
--- a/src/ZF/Trancl.ML	Thu Apr 10 09:08:05 1997 +0200
+++ b/src/ZF/Trancl.ML	Thu Apr 10 10:55:37 1997 +0200
@@ -3,7 +3,7 @@
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 
-For trancl.thy.  Transitive closure of a relation
+Transitive closure of a relation
 *)
 
 open Trancl;
@@ -11,7 +11,7 @@
 goal Trancl.thy "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))";
 by (rtac bnd_monoI 1);
 by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2));
-by (Fast_tac 1);
+by (Blast_tac 1);
 qed "rtrancl_bnd_mono";
 
 val [prem] = goalw Trancl.thy [rtrancl_def] "r<=s ==> r^* <= s^*";
@@ -51,11 +51,11 @@
 (*The premise ensures that r consists entirely of pairs*)
 val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^*";
 by (cut_facts_tac prems 1);
-by (fast_tac (!claset addIs [r_into_rtrancl]) 1);
+by (blast_tac (!claset addIs [r_into_rtrancl]) 1);
 qed "r_subset_rtrancl";
 
 goal Trancl.thy "field(r^*) = field(r)";
-by (fast_tac (!claset addIs [r_into_rtrancl] 
+by (blast_tac (!claset addIs [r_into_rtrancl] 
                     addSDs [rtrancl_type RS subsetD]) 1);
 qed "rtrancl_field";
 
@@ -131,7 +131,7 @@
 (*The premise ensures that r consists entirely of pairs*)
 val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^+";
 by (cut_facts_tac prems 1);
-by (fast_tac (!claset addIs [r_into_trancl]) 1);
+by (blast_tac (!claset addIs [r_into_trancl]) 1);
 qed "r_subset_trancl";
 
 (*intro rule by definition: from r^* and r  *)
@@ -159,7 +159,7 @@
 (*by induction on this formula*)
 by (subgoal_tac "ALL z. <y,z> : r --> P(z)" 1);
 (*now solve first subgoal: this formula is sufficient*)
-by (Fast_tac 1);
+by (Blast_tac 1);
 by (etac rtrancl_induct 1);
 by (ALLGOALS (fast_tac (!claset addIs (rtrancl_into_trancl1::prems))));
 qed "trancl_induct";
@@ -174,11 +174,11 @@
 by (fast_tac (!claset addIs prems) 1);
 by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1);
 by (etac rtranclE 1);
-by (ALLGOALS (fast_tac (!claset addIs [rtrancl_into_trancl1])));
+by (ALLGOALS (blast_tac (!claset addIs [rtrancl_into_trancl1])));
 qed "tranclE";
 
 goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)";
-by (fast_tac (!claset addEs [rtrancl_type RS subsetD RS SigmaE2]) 1);
+by (blast_tac (!claset addEs [rtrancl_type RS subsetD RS SigmaE2]) 1);
 qed "trancl_type";
 
 val [prem] = goalw Trancl.thy [trancl_def] "r<=s ==> r^+ <= s^+";