hide constant Quickcheck.random
authorhaftmann
Mon, 15 Jun 2009 16:13:03 +0200
changeset 31641 feea4d3d743d
parent 31640 51fb047168b7
child 31642 ce68241f711f
hide constant Quickcheck.random
src/HOL/Library/Fin_Fun.thy
src/HOL/Quickcheck.thy
src/HOL/Rational.thy
src/HOL/RealDef.thy
src/HOL/Tools/quickcheck_generators.ML
--- a/src/HOL/Library/Fin_Fun.thy	Mon Jun 15 11:01:18 2009 +0200
+++ b/src/HOL/Library/Fin_Fun.thy	Mon Jun 15 16:13:03 2009 +0200
@@ -329,7 +329,7 @@
          (1, random j o\<rightarrow> (\<lambda>y. Pair (valtermify_finfun_const y)))])"
 
 definition 
-  "random i = random_finfun_aux i i"
+  "Quickcheck.random i = random_finfun_aux i i"
 
 instance ..
 
--- a/src/HOL/Quickcheck.thy	Mon Jun 15 11:01:18 2009 +0200
+++ b/src/HOL/Quickcheck.thy	Mon Jun 15 16:13:03 2009 +0200
@@ -137,7 +137,7 @@
 code_reserved Quickcheck Quickcheck_Generators
 
 
-hide (open) const collapse beyond random_fun_aux random_fun_lift
+hide (open) const random collapse beyond random_fun_aux random_fun_lift
 
 no_notation fcomp (infixl "o>" 60)
 no_notation scomp (infixl "o\<rightarrow>" 60)
--- a/src/HOL/Rational.thy	Mon Jun 15 11:01:18 2009 +0200
+++ b/src/HOL/Rational.thy	Mon Jun 15 16:13:03 2009 +0200
@@ -1012,7 +1012,7 @@
 begin
 
 definition
-  "random i = random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
+  "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
      let j = Code_Numeral.int_of (denom + 1)
      in valterm_fract num (j, \<lambda>u. Code_Eval.term_of j))))"
 
--- a/src/HOL/RealDef.thy	Mon Jun 15 11:01:18 2009 +0200
+++ b/src/HOL/RealDef.thy	Mon Jun 15 16:13:03 2009 +0200
@@ -1137,7 +1137,7 @@
 begin
 
 definition
-  "random i = random i o\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
+  "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
 
 instance ..
 
--- a/src/HOL/Tools/quickcheck_generators.ML	Mon Jun 15 11:01:18 2009 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML	Mon Jun 15 16:13:03 2009 +0200
@@ -49,7 +49,7 @@
       mk_scomp (mk_termtyp ty) @{typ "term list option"} @{typ Random.seed} t
         (mk_split ty $ Abs ("", ty, Abs ("", @{typ "unit => term"}, t')));
     fun mk_bindclause (_, _, i, ty) = mk_scomp_split ty
-      (Sign.mk_const thy (@{const_name random}, [ty]) $ Bound i);
+      (Sign.mk_const thy (@{const_name Quickcheck.random}, [ty]) $ Bound i);
   in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end;
 
 fun compile_generator_expr thy t =
@@ -98,7 +98,7 @@
     val Ttm = mk_termifyT T;
     val typtm = mk_termifyT typ;
     fun mk_const c Ts = Const (c, Sign.const_instance thy (c, Ts));
-    fun mk_random T = mk_const @{const_name random} [T];
+    fun mk_random T = mk_const @{const_name Quickcheck.random} [T];
     val size = @{term "j::code_numeral"};
     val v = "x";
     val t_v = Free (v, typtm);