use "get_goal" rather than "flat_goal" in Auto Quickcheck, since we don't need the extra facts for counterexample generation
--- a/src/Tools/quickcheck.ML Tue Oct 27 21:53:13 2009 +0100
+++ b/src/Tools/quickcheck.ML Wed Oct 28 11:55:48 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 (_, (_, st)) = Proof.get_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