--- a/src/Pure/Isar/object_logic.ML Thu Jan 17 21:05:40 2002 +0100
+++ b/src/Pure/Isar/object_logic.ML Thu Jan 17 21:05:58 2002 +0100
@@ -16,7 +16,7 @@
val fixed_judgment: Sign.sg -> string -> term
val declare_atomize: theory attribute
val declare_rulify: theory attribute
- val atomize_cterm: cterm -> cterm
+ val atomize_term: Sign.sg -> term -> term
val atomize_tac: int -> tactic
val atomize_goal: int -> thm -> thm
val rulify: thm -> thm
@@ -127,12 +127,11 @@
(MetaSimplifier.forall_conv
(MetaSimplifier.goals_conv (K true) (Tactic.rewrite true rews)))));
-fun atomize_cterm ct =
- let val sg = #sign (Thm.rep_cterm ct)
- in Thm.cterm_fun (drop_judgment sg) (Tactic.rewrite_cterm true (get_atomize sg) ct) end;
+fun atomize_term sg =
+ drop_judgment sg o MetaSimplifier.rewrite_term sg (get_atomize sg);
fun atomize_tac i st =
- if Logic.has_meta_prems (#prop (Thm.rep_thm st)) i then
+ if Logic.has_meta_prems (Thm.prop_of st) i then
(rewrite_prems_tac (get_atomize (Thm.sign_of_thm st)) i) st
else all_tac st;