src/Tools/quickcheck.ML
changeset 33291 93f0238151f6
parent 32966 5b21661fe618
child 33522 737589bb9bb8
child 33566 1c62ac4ef6d1
--- a/src/Tools/quickcheck.ML	Wed Oct 28 22:02:53 2009 +0100
+++ b/src/Tools/quickcheck.ML	Wed Oct 28 22:04:57 2009 +0100
@@ -143,7 +143,7 @@
     val thy = Proof.theory_of state;
     fun strip (Const ("all", _) $ Abs (_, _, t)) = strip t
       | strip t = t;
-    val (_, st) = Proof.flat_goal state;
+    val {goal = st, ...} = Proof.raw_goal state;
     val (gi, frees) = Logic.goal_params (prop_of st) i;
     val gi' = Logic.list_implies (assms, subst_bounds (frees, strip gi))
       |> monomorphic_term thy insts default_T