src/FOL/hypsubstdata.ML
changeset 12353 e4be26df707a
parent 12345 5d1436d1c268
child 20974 2a29a21825d1
--- a/src/FOL/hypsubstdata.ML	Tue Dec 04 02:00:14 2001 +0100
+++ b/src/FOL/hypsubstdata.ML	Tue Dec 04 02:00:45 2001 +0100
@@ -14,12 +14,8 @@
   val subst = subst
   val sym = sym
   val thin_refl = prove_goal (the_context ())
-		  "!!X. [|x=x; PROP W|] ==> PROP W" (K [atac 1]);
+                  "!!X. [|x=x; PROP W|] ==> PROP W" (K [atac 1]);
   end;
 
 structure Hypsubst = HypsubstFun(Hypsubst_Data);
 open Hypsubst;
-
-fun hyp_subst_tac' i thm = if i <= Thm.nprems_of thm andalso
-  Term.exists_Const (fn ("op =", Type (_, [T, _])) => T <> Type ("o", []) | _ => false)
-    (Library.nth_elem (i - 1, Thm.prems_of thm)) then hyp_subst_tac i thm else no_tac thm;