--- a/src/HOL/Quickcheck_Narrowing.thy Mon Jul 18 10:34:21 2011 +0200
+++ b/src/HOL/Quickcheck_Narrowing.thy Mon Jul 18 10:34:21 2011 +0200
@@ -431,9 +431,8 @@
*)
hide_type code_int narrowing_type narrowing_term cons property
-hide_const int_of of_int nth error toEnum marker empty
- C cons conv nonEmpty "apply" sum ensure_testable all exists
-hide_const (open) Var Ctr
+hide_const int_of of_int nth error toEnum marker empty C conv nonEmpty ensure_testable all exists
+hide_const (open) Var Ctr "apply" sum cons
hide_fact empty_def cons_def conv.simps nonEmpty.simps apply_def sum_def ensure_testable_def all_def exists_def
--- a/src/HOL/Rat.thy Mon Jul 18 10:34:21 2011 +0200
+++ b/src/HOL/Rat.thy Mon Jul 18 10:34:21 2011 +0200
@@ -1171,6 +1171,17 @@
end
+instantiation rat :: narrowing
+begin
+
+definition
+ "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Fract) narrowing) narrowing"
+
+instance ..
+
+end
+
+
text {* Setup for SML code generator *}
types_code
--- a/src/HOL/RealDef.thy Mon Jul 18 10:34:21 2011 +0200
+++ b/src/HOL/RealDef.thy Mon Jul 18 10:34:21 2011 +0200
@@ -1805,6 +1805,17 @@
end
+instantiation real :: narrowing
+begin
+
+definition
+ "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing"
+
+instance ..
+
+end
+
+
text {* Setup for SML code generator *}
types_code