src/HOL/Library/DAList.thy
changeset 72607 feebdaa346e5
parent 72581 de581f98a3a1
--- a/src/HOL/Library/DAList.thy	Sun Nov 15 07:17:05 2020 +0000
+++ b/src/HOL/Library/DAList.thy	Sun Nov 15 07:17:06 2020 +0000
@@ -128,21 +128,21 @@
 
 subsection \<open>Quickcheck generators\<close>
 
-definition (in term_syntax)
+context
+  includes state_combinator_syntax term_syntax
+begin
+
+definition
   valterm_empty :: "('key :: typerep, 'value :: typerep) alist \<times> (unit \<Rightarrow> Code_Evaluation.term)"
   where "valterm_empty = Code_Evaluation.valtermify empty"
 
-definition (in term_syntax)
+definition
   valterm_update :: "'key :: typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow>
   'value :: typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow>
   ('key, 'value) alist \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow>
   ('key, 'value) alist \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
   [code_unfold]: "valterm_update k v a = Code_Evaluation.valtermify update {\<cdot>} k {\<cdot>} v {\<cdot>}a"
 
-context
-  includes state_combinator_syntax
-begin
-
 fun random_aux_alist
 where
   "random_aux_alist i j =