--- a/src/FOL/blastdata.ML Tue Jan 16 00:27:37 2001 +0100
+++ b/src/FOL/blastdata.ML Tue Jan 16 00:28:50 2001 +0100
@@ -10,7 +10,7 @@
val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
val claset = Cla.claset
val rep_cs = Cla.rep_cs
- val atomize = thms "atomize'"
+ val atomize = atomize_rules
val cla_modifiers = Cla.cla_modifiers;
val cla_meth' = Cla.cla_meth'
end;
--- a/src/FOL/cladata.ML Tue Jan 16 00:27:37 2001 +0100
+++ b/src/FOL/cladata.ML Tue Jan 16 00:28:50 2001 +0100
@@ -15,6 +15,10 @@
compatibliity with strange things done in many existing proofs *)
val cla_make_elim = Make_Elim.make_elim;
+val atomize_rules = thms "atomize'";
+val atomize_tac = Method.atomize_tac atomize_rules;
+val atomize_strip_tac = Method.atomize_strip_tac (atomize_rules, [impI, allI]);
+
(*** Applying ClassicalFun to create a classical prover ***)
structure Classical_Data =
struct
@@ -24,7 +28,7 @@
val classical = classical
val sizef = size_of_thm
val hyp_subst_tacs=[hyp_subst_tac]
- val atomize = thms "atomize'"
+ val atomize = atomize_rules
end;
structure Cla = ClassicalFun(Classical_Data);
--- a/src/HOL/arith_data.ML Tue Jan 16 00:27:37 2001 +0100
+++ b/src/HOL/arith_data.ML Tue Jan 16 00:28:50 2001 +0100
@@ -407,8 +407,6 @@
*)
local
-val atomize_tac = Method.atomize_tac (thms "atomize'");
-
fun raw_arith_tac i st =
refute_tac (K true) (REPEAT o split_tac (#splits (ArithTheoryData.get_sg (Thm.sign_of_thm st))))
((REPEAT_DETERM o etac linorder_neqE) THEN' fast_arith_tac) i st;
--- a/src/HOL/blastdata.ML Tue Jan 16 00:27:37 2001 +0100
+++ b/src/HOL/blastdata.ML Tue Jan 16 00:28:50 2001 +0100
@@ -23,7 +23,7 @@
val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
val claset = Classical.claset
val rep_cs = Classical.rep_cs
- val atomize = thms "atomize'"
+ val atomize = atomize_rules
val cla_modifiers = Classical.cla_modifiers;
val cla_meth' = Classical.cla_meth'
end;
--- a/src/HOL/cladata.ML Tue Jan 16 00:27:37 2001 +0100
+++ b/src/HOL/cladata.ML Tue Jan 16 00:28:50 2001 +0100
@@ -38,6 +38,10 @@
compatibliity with strange things done in many existing proofs *)
val cla_make_elim = Make_Elim.make_elim;
+val atomize_rules = thms "atomize'";
+val atomize_tac = Method.atomize_tac atomize_rules;
+val atomize_strip_tac = Method.atomize_strip_tac (atomize_rules, [impI, allI]);
+
(*** Applying ClassicalFun to create a classical prover ***)
structure Classical_Data =
struct
@@ -47,7 +51,7 @@
val classical = classical
val sizef = size_of_thm
val hyp_subst_tacs=[hyp_subst_tac]
- val atomize = thms "atomize'"
+ val atomize = atomize_rules
end;
structure Classical = ClassicalFun(Classical_Data);