src/CCL/Wfd.thy
changeset 32211 752dbe71cc89
parent 31902 862ae16a799d
child 32282 853ef99c04cc
--- a/src/CCL/Wfd.thy	Sun Jul 26 19:38:02 2009 +0200
+++ b/src/CCL/Wfd.thy	Sun Jul 26 19:54:37 2009 +0200
@@ -498,14 +498,13 @@
   structure Data = Named_Thms(val name = "eval" val description = "evaluation rules");
 in
 
-fun eval_tac ctxt ths =
-  METAHYPS (fn prems =>
-    DEPTH_SOLVE_1 (resolve_tac (ths @ prems @ Data.get ctxt) 1)) 1;
+fun eval_tac ctxt ths i =
+  DEPTH_SOLVE_1 (resolve_tac (ths @ Data.get ctxt) i ORELSE assume_tac i);
 
 val eval_setup =
   Data.setup #>
   Method.setup @{binding eval}
-    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (CHANGED (eval_tac ctxt ths))))
+    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED o eval_tac ctxt ths)))
     "evaluation";
 
 end;