src/HOL/Rat.thy
changeset 72607 feebdaa346e5
parent 72581 de581f98a3a1
child 73109 783406dd051e
--- a/src/HOL/Rat.thy	Sun Nov 15 07:17:05 2020 +0000
+++ b/src/HOL/Rat.thy	Sun Nov 15 07:17:06 2020 +0000
@@ -998,12 +998,18 @@
 
 text \<open>Quickcheck\<close>
 
-definition (in term_syntax)
+context
+  includes term_syntax
+begin
+
+definition
   valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow>
     int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow>
     rat \<times> (unit \<Rightarrow> Code_Evaluation.term)"
   where [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\<cdot>} k {\<cdot>} l"
 
+end
+
 instantiation rat :: random
 begin