src/HOL/Library/Eval_Witness.thy
changeset 27103 d8549f4d900b
parent 26939 1035c89b4c02
child 27368 9f90ac19e32b
--- a/src/HOL/Library/Eval_Witness.thy	Tue Jun 10 15:30:01 2008 +0200
+++ b/src/HOL/Library/Eval_Witness.thy	Tue Jun 10 15:30:06 2008 +0200
@@ -46,6 +46,15 @@
 instance bool :: ml_equiv ..
 instance list :: (ml_equiv) ml_equiv ..
 
+ML {*
+structure Eval_Witness_Method =
+struct
+
+val eval_ref : (unit -> bool) option ref = ref NONE;
+
+end;
+*}
+
 oracle eval_witness_oracle ("term * string list") = {* fn thy => fn (goal, ws) => 
 let
   fun check_type T = 
@@ -59,7 +68,7 @@
     | dest_exs _ _ = sys_error "dest_exs";
   val t = dest_exs (length ws) (HOLogic.dest_Trueprop goal);
 in
-  if CodePackage.satisfies thy t ws
+  if CodeTarget.eval_term ("Eval_Witness_Method.eval_ref", Eval_Witness_Method.eval_ref) thy t ws
   then goal
   else HOLogic.Trueprop $ HOLogic.true_const (*dummy*)
 end