src/HOL/Real.thy
changeset 72607 feebdaa346e5
parent 72581 de581f98a3a1
child 73932 fd21b4a93043
--- a/src/HOL/Real.thy	Sun Nov 15 07:17:05 2020 +0000
+++ b/src/HOL/Real.thy	Sun Nov 15 07:17:06 2020 +0000
@@ -1634,10 +1634,16 @@
 
 text \<open>Quickcheck\<close>
 
-definition (in term_syntax)
+context
+  includes term_syntax
+begin
+
+definition
   valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)"
   where [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
 
+end
+
 instantiation real :: random
 begin