src/HOL/Tools/Quickcheck/exhaustive_generators.ML
changeset 67149 e61557884799
parent 62981 3a01f1f58630
child 70308 7f568724d67e
--- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Wed Dec 06 19:34:59 2017 +0100
+++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Wed Dec 06 20:43:09 2017 +0100
@@ -38,30 +38,30 @@
 
 (** dynamic options **)
 
-val smart_quantifier = Attrib.setup_config_bool @{binding quickcheck_smart_quantifier} (K true)
-val optimise_equality = Attrib.setup_config_bool @{binding quickcheck_optimise_equality} (K true)
+val smart_quantifier = Attrib.setup_config_bool \<^binding>\<open>quickcheck_smart_quantifier\<close> (K true)
+val optimise_equality = Attrib.setup_config_bool \<^binding>\<open>quickcheck_optimise_equality\<close> (K true)
 
-val fast = Attrib.setup_config_bool @{binding quickcheck_fast} (K false)
-val bounded_forall = Attrib.setup_config_bool @{binding quickcheck_bounded_forall} (K false)
-val full_support = Attrib.setup_config_bool @{binding quickcheck_full_support} (K true)
-val quickcheck_pretty = Attrib.setup_config_bool @{binding quickcheck_pretty} (K true)
+val fast = Attrib.setup_config_bool \<^binding>\<open>quickcheck_fast\<close> (K false)
+val bounded_forall = Attrib.setup_config_bool \<^binding>\<open>quickcheck_bounded_forall\<close> (K false)
+val full_support = Attrib.setup_config_bool \<^binding>\<open>quickcheck_full_support\<close> (K true)
+val quickcheck_pretty = Attrib.setup_config_bool \<^binding>\<open>quickcheck_pretty\<close> (K true)
 
 
 (** abstract syntax **)
 
-fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit \<Rightarrow> Code_Evaluation.term"})
+fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\<open>unit \<Rightarrow> Code_Evaluation.term\<close>)
 
-val size = @{term "i :: natural"}
-val size_pred = @{term "(i :: natural) - 1"}
-val size_ge_zero = @{term "(i :: natural) > 0"}
+val size = \<^term>\<open>i :: natural\<close>
+val size_pred = \<^term>\<open>(i :: natural) - 1\<close>
+val size_ge_zero = \<^term>\<open>(i :: natural) > 0\<close>
 
 fun mk_none_continuation (x, y) =
-  let val (T as Type (@{type_name option}, _)) = fastype_of x
-  in Const (@{const_name orelse}, T --> T --> T) $ x $ y end
+  let val (T as Type (\<^type_name>\<open>option\<close>, _)) = fastype_of x
+  in Const (\<^const_name>\<open>orelse\<close>, T --> T --> T) $ x $ y end
 
 fun mk_if (b, t, e) =
   let val T = fastype_of t
-  in Const (@{const_name If}, @{typ bool} --> T --> T --> T) $ b $ t $ e end
+  in Const (\<^const_name>\<open>If\<close>, \<^typ>\<open>bool\<close> --> T --> T --> T) $ b $ t $ e end
 
 
 (* handling inductive datatypes *)
@@ -72,19 +72,19 @@
 
 exception Counterexample of term list
 
-val resultT =  @{typ "(bool * term list) option"}
+val resultT = \<^typ>\<open>(bool \<times> term list) option\<close>
 
 val exhaustiveN = "exhaustive"
 val full_exhaustiveN = "full_exhaustive"
 val bounded_forallN = "bounded_forall"
 
-fun fast_exhaustiveT T = (T --> @{typ unit}) --> @{typ natural} --> @{typ unit}
+fun fast_exhaustiveT T = (T --> \<^typ>\<open>unit\<close>) --> \<^typ>\<open>natural\<close> --> \<^typ>\<open>unit\<close>
 
-fun exhaustiveT T = (T --> resultT) --> @{typ natural} --> resultT
+fun exhaustiveT T = (T --> resultT) --> \<^typ>\<open>natural\<close> --> resultT
 
-fun bounded_forallT T = (T --> @{typ bool}) --> @{typ natural} --> @{typ bool}
+fun bounded_forallT T = (T --> \<^typ>\<open>bool\<close>) --> \<^typ>\<open>natural\<close> --> \<^typ>\<open>bool\<close>
 
-fun full_exhaustiveT T = (termifyT T --> resultT) --> @{typ natural} --> resultT
+fun full_exhaustiveT T = (termifyT T --> resultT) --> \<^typ>\<open>natural\<close> --> resultT
 
 fun check_allT T = (termifyT T --> resultT) --> resultT
 
@@ -119,35 +119,35 @@
 fun mk_equations functerms =
   let
     fun test_function T = Free ("f", T --> resultT)
-    val mk_call = gen_mk_call (fn T => Const (@{const_name exhaustive}, exhaustiveT T))
+    val mk_call = gen_mk_call (fn T => Const (\<^const_name>\<open>exhaustive\<close>, exhaustiveT T))
     val mk_aux_call = gen_mk_aux_call functerms
     val mk_consexpr = gen_mk_consexpr test_function
     fun mk_rhs exprs =
-      mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (@{const_name None}, resultT))
+      mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (\<^const_name>\<open>None\<close>, resultT))
   in mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) end
 
 fun mk_bounded_forall_equations functerms =
   let
-    fun test_function T = Free ("P", T --> @{typ bool})
-    val mk_call = gen_mk_call (fn T => Const (@{const_name bounded_forall}, bounded_forallT T))
+    fun test_function T = Free ("P", T --> \<^typ>\<open>bool\<close>)
+    val mk_call = gen_mk_call (fn T => Const (\<^const_name>\<open>bounded_forall\<close>, bounded_forallT T))
     val mk_aux_call = gen_mk_aux_call functerms
     val mk_consexpr = gen_mk_consexpr test_function
-    fun mk_rhs exprs = mk_if (size_ge_zero, foldr1 HOLogic.mk_conj exprs, @{term True})
+    fun mk_rhs exprs = mk_if (size_ge_zero, foldr1 HOLogic.mk_conj exprs, \<^term>\<open>True\<close>)
   in mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) end
 
 fun mk_full_equations functerms =
   let
     fun test_function T = Free ("f", termifyT T --> resultT)
     fun case_prod_const T =
-      HOLogic.case_prod_const (T, @{typ "unit \<Rightarrow> Code_Evaluation.term"}, resultT)
+      HOLogic.case_prod_const (T, \<^typ>\<open>unit \<Rightarrow> Code_Evaluation.term\<close>, resultT)
     fun mk_call T =
       let
-        val full_exhaustive = Const (@{const_name full_exhaustive}, full_exhaustiveT T)
+        val full_exhaustive = Const (\<^const_name>\<open>full_exhaustive\<close>, full_exhaustiveT T)
       in
         (T,
           fn t =>
             full_exhaustive $
-              (case_prod_const T $ absdummy T (absdummy @{typ "unit \<Rightarrow> Code_Evaluation.term"} t)) $
+              (case_prod_const T $ absdummy T (absdummy \<^typ>\<open>unit \<Rightarrow> Code_Evaluation.term\<close> t)) $
               size_pred)
       end
     fun mk_aux_call fTs (k, _) (tyco, Ts) =
@@ -158,7 +158,7 @@
         (T,
           fn t =>
             nth functerms k $
-              (case_prod_const T $ absdummy T (absdummy @{typ "unit \<Rightarrow> Code_Evaluation.term"} t)) $
+              (case_prod_const T $ absdummy T (absdummy \<^typ>\<open>unit \<Rightarrow> Code_Evaluation.term\<close> t)) $
               size_pred)
       end
     fun mk_consexpr simpleT (c, xs) =
@@ -167,21 +167,21 @@
         val constr = Const (c, Ts ---> simpleT)
         val bounds = map (fn x => Bound (2 * x + 1)) (((length xs) - 1) downto 0)
         val Eval_App =
-          Const (@{const_name Code_Evaluation.App},
+          Const (\<^const_name>\<open>Code_Evaluation.App\<close>,
             HOLogic.termT --> HOLogic.termT --> HOLogic.termT)
         val Eval_Const =
-          Const (@{const_name Code_Evaluation.Const},
-            HOLogic.literalT --> @{typ typerep} --> HOLogic.termT)
+          Const (\<^const_name>\<open>Code_Evaluation.Const\<close>,
+            HOLogic.literalT --> \<^typ>\<open>typerep\<close> --> HOLogic.termT)
         val term =
-          fold (fn u => fn t => Eval_App $ t $ (u $ @{term "()"}))
+          fold (fn u => fn t => Eval_App $ t $ (u $ \<^term>\<open>()\<close>))
             bounds (Eval_Const $ HOLogic.mk_literal c $ HOLogic.mk_typerep (Ts ---> simpleT))
         val start_term =
           test_function simpleT $
-            (HOLogic.pair_const simpleT @{typ "unit \<Rightarrow> Code_Evaluation.term"} $
-              (list_comb (constr, bounds)) $ absdummy @{typ unit} term)
+            (HOLogic.pair_const simpleT \<^typ>\<open>unit \<Rightarrow> Code_Evaluation.term\<close> $
+              (list_comb (constr, bounds)) $ absdummy \<^typ>\<open>unit\<close> term)
       in fold_rev (fn f => fn t => f t) fns start_term end
     fun mk_rhs exprs =
-      mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (@{const_name None}, resultT))
+      mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (\<^const_name>\<open>None\<close>, resultT))
   in mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) end
 
 
@@ -212,17 +212,17 @@
 
 val instantiate_bounded_forall_datatype =
   instantiate_datatype
-    ("bounded universal quantifiers", bounded_forallN, @{sort bounded_forall},
+    ("bounded universal quantifiers", bounded_forallN, \<^sort>\<open>bounded_forall\<close>,
       mk_bounded_forall_equations, bounded_forallT, ["P", "i"])
 
 val instantiate_exhaustive_datatype =
   instantiate_datatype
-    ("exhaustive generators", exhaustiveN, @{sort exhaustive},
+    ("exhaustive generators", exhaustiveN, \<^sort>\<open>exhaustive\<close>,
       mk_equations, exhaustiveT, ["f", "i"])
 
 val instantiate_full_exhaustive_datatype =
   instantiate_datatype
-    ("full exhaustive generators", full_exhaustiveN, @{sort full_exhaustive},
+    ("full exhaustive generators", full_exhaustiveN, \<^sort>\<open>full_exhaustive\<close>,
       mk_full_equations, full_exhaustiveT, ["f", "i"])
 
 
@@ -230,15 +230,15 @@
 
 fun mk_let_expr (x, t, e) genuine =
   let val (T1, T2) = (fastype_of x, fastype_of (e genuine))
-  in Const (@{const_name Let}, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine) end
+  in Const (\<^const_name>\<open>Let\<close>, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine) end
 
 fun mk_safe_let_expr genuine_only none safe (x, t, e) genuine =
   let
     val (T1, T2) = (fastype_of x, fastype_of (e genuine))
-    val if_t = Const (@{const_name If}, @{typ bool} --> T2 --> T2 --> T2)
+    val if_t = Const (\<^const_name>\<open>If\<close>, \<^typ>\<open>bool\<close> --> T2 --> T2 --> T2)
   in
-    Const (@{const_name Quickcheck_Random.catch_match}, T2 --> T2 --> T2) $
-      (Const (@{const_name Let}, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine)) $
+    Const (\<^const_name>\<open>Quickcheck_Random.catch_match\<close>, T2 --> T2 --> T2) $
+      (Const (\<^const_name>\<open>Let\<close>, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine)) $
       (if_t $ genuine_only $ none $ safe false)
   end
 
@@ -323,30 +323,30 @@
     val frees = map Free (Term.add_frees t [])
     fun lookup v = the (AList.lookup (op =) (names ~~ frees) v)
     val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt'
-    val depth = Free (depth_name, @{typ natural})
+    val depth = Free (depth_name, \<^typ>\<open>natural\<close>)
     fun return _ =
-      @{term "throw_Counterexample :: term list \<Rightarrow> unit"} $
-        (HOLogic.mk_list @{typ term}
+      \<^term>\<open>throw_Counterexample :: term list \<Rightarrow> unit\<close> $
+        (HOLogic.mk_list \<^typ>\<open>term\<close>
           (map (fn t => HOLogic.mk_term_of (fastype_of t) t) (frees @ eval_terms)))
     fun mk_exhaustive_closure (free as Free (_, T)) t =
-      Const (@{const_name fast_exhaustive}, fast_exhaustiveT T) $ lambda free t $ depth
-    val none_t = @{term "()"}
+      Const (\<^const_name>\<open>fast_exhaustive\<close>, fast_exhaustiveT T) $ lambda free t $ depth
+    val none_t = \<^term>\<open>()\<close>
     fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine)
     fun mk_let _ def v_opt t e = mk_let_expr (the_default def v_opt, t, e)
     val mk_test_term =
       mk_test_term lookup mk_exhaustive_closure mk_safe_if mk_let none_t return ctxt
-  in lambda depth (@{term "catch_Counterexample :: unit => term list option"} $ mk_test_term t) end
+  in lambda depth (\<^term>\<open>catch_Counterexample :: unit => term list option\<close> $ mk_test_term t) end
 
 fun mk_unknown_term T =
-  HOLogic.reflect_term (Const (@{const_name unknown}, T))
+  HOLogic.reflect_term (Const (\<^const_name>\<open>unknown\<close>, T))
 
 fun mk_safe_term t =
-  @{term "Quickcheck_Random.catch_match :: term \<Rightarrow> term \<Rightarrow> term"} $
+  \<^term>\<open>Quickcheck_Random.catch_match :: term \<Rightarrow> term \<Rightarrow> term\<close> $
     (HOLogic.mk_term_of (fastype_of t) t) $ mk_unknown_term (fastype_of t)
 
 fun mk_return t genuine =
-  @{term "Some :: bool \<times> term list \<Rightarrow> (bool \<times> term list) option"} $
-    (HOLogic.pair_const @{typ bool} @{typ "term list"} $
+  \<^term>\<open>Some :: bool \<times> term list \<Rightarrow> (bool \<times> term list) option\<close> $
+    (HOLogic.pair_const \<^typ>\<open>bool\<close> \<^typ>\<open>term list\<close> $
       Quickcheck_Common.reflect_bool genuine $ t)
 
 fun mk_generator_expr ctxt (t, eval_terms) =
@@ -357,14 +357,14 @@
     fun lookup v = the (AList.lookup (op =) (names ~~ frees) v)
     val ([depth_name, genuine_only_name], _) =
       Variable.variant_fixes ["depth", "genuine_only"] ctxt'
-    val depth = Free (depth_name, @{typ natural})
-    val genuine_only = Free (genuine_only_name, @{typ bool})
+    val depth = Free (depth_name, \<^typ>\<open>natural\<close>)
+    val genuine_only = Free (genuine_only_name, \<^typ>\<open>bool\<close>)
     val return =
-      mk_return (HOLogic.mk_list @{typ term}
+      mk_return (HOLogic.mk_list \<^typ>\<open>term\<close>
         (map (fn t => HOLogic.mk_term_of (fastype_of t) t) frees @ map mk_safe_term eval_terms))
     fun mk_exhaustive_closure (free as Free (_, T)) t =
-      Const (@{const_name exhaustive}, exhaustiveT T) $ lambda free t $ depth
-    val none_t = Const (@{const_name None}, resultT)
+      Const (\<^const_name>\<open>exhaustive\<close>, exhaustiveT T) $ lambda free t $ depth
+    val none_t = Const (\<^const_name>\<open>None\<close>, resultT)
     val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t
     fun mk_let safe def v_opt t e =
       mk_safe_let_expr genuine_only none_t safe (the_default def v_opt, t, e)
@@ -380,28 +380,28 @@
     val ([depth_name, genuine_only_name], ctxt'') =
       Variable.variant_fixes ["depth", "genuine_only"] ctxt'
     val (term_names, _) = Variable.variant_fixes (map (prefix "t_") names) ctxt''
-    val depth = Free (depth_name, @{typ natural})
-    val genuine_only = Free (genuine_only_name, @{typ bool})
-    val term_vars = map (fn n => Free (n, @{typ "unit \<Rightarrow> term"})) term_names
+    val depth = Free (depth_name, \<^typ>\<open>natural\<close>)
+    val genuine_only = Free (genuine_only_name, \<^typ>\<open>bool\<close>)
+    val term_vars = map (fn n => Free (n, \<^typ>\<open>unit \<Rightarrow> term\<close>)) term_names
     fun lookup v = the (AList.lookup (op =) (names ~~ (frees ~~ term_vars)) v)
     val return =
       mk_return
-        (HOLogic.mk_list @{typ term}
-          (map (fn v => v $ @{term "()"}) term_vars @ map mk_safe_term eval_terms))
+        (HOLogic.mk_list \<^typ>\<open>term\<close>
+          (map (fn v => v $ \<^term>\<open>()\<close>) term_vars @ map mk_safe_term eval_terms))
     fun mk_exhaustive_closure (free as Free (_, T), term_var) t =
-      if Sign.of_sort thy (T, @{sort check_all}) then
-        Const (@{const_name check_all}, check_allT T) $
-          (HOLogic.case_prod_const (T, @{typ "unit \<Rightarrow> term"}, resultT) $
+      if Sign.of_sort thy (T, \<^sort>\<open>check_all\<close>) then
+        Const (\<^const_name>\<open>check_all\<close>, check_allT T) $
+          (HOLogic.case_prod_const (T, \<^typ>\<open>unit \<Rightarrow> term\<close>, resultT) $
             lambda free (lambda term_var t))
       else
-        Const (@{const_name full_exhaustive}, full_exhaustiveT T) $
-          (HOLogic.case_prod_const (T, @{typ "unit \<Rightarrow> term"}, resultT) $
+        Const (\<^const_name>\<open>full_exhaustive\<close>, full_exhaustiveT T) $
+          (HOLogic.case_prod_const (T, \<^typ>\<open>unit \<Rightarrow> term\<close>, resultT) $
             lambda free (lambda term_var t)) $ depth
-    val none_t = Const (@{const_name None}, resultT)
+    val none_t = Const (\<^const_name>\<open>None\<close>, resultT)
     val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t
     fun mk_let safe _ (SOME (v, term_var)) t e =
           mk_safe_let_expr genuine_only none_t safe (v, t,
-            e #> subst_free [(term_var, absdummy @{typ unit} (mk_safe_term t))])
+            e #> subst_free [(term_var, absdummy \<^typ>\<open>unit\<close> (mk_safe_term t))])
       | mk_let safe v NONE t e = mk_safe_let_expr genuine_only none_t safe (v, t, e)
     val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_if mk_let none_t return ctxt
   in lambda genuine_only (lambda depth (mk_test_term t)) end
@@ -409,37 +409,37 @@
 fun mk_parametric_generator_expr mk_generator_expr =
   Quickcheck_Common.gen_mk_parametric_generator_expr
     ((mk_generator_expr,
-      absdummy @{typ bool} (absdummy @{typ natural} (Const (@{const_name None}, resultT)))),
-      @{typ bool} --> @{typ natural} --> resultT)
+      absdummy \<^typ>\<open>bool\<close> (absdummy \<^typ>\<open>natural\<close> (Const (\<^const_name>\<open>None\<close>, resultT)))),
+      \<^typ>\<open>bool\<close> --> \<^typ>\<open>natural\<close> --> resultT)
 
 fun mk_validator_expr ctxt t =
   let
-    fun bounded_forallT T = (T --> @{typ bool}) --> @{typ natural} --> @{typ bool}
+    fun bounded_forallT T = (T --> \<^typ>\<open>bool\<close>) --> \<^typ>\<open>natural\<close> --> \<^typ>\<open>bool\<close>
     val ctxt' = Variable.auto_fixes t ctxt
     val names = Term.add_free_names t []
     val frees = map Free (Term.add_frees t [])
     fun lookup v = the (AList.lookup (op =) (names ~~ frees) v)
     val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt'
-    val depth = Free (depth_name, @{typ natural})
+    val depth = Free (depth_name, \<^typ>\<open>natural\<close>)
     fun mk_bounded_forall (Free (s, T)) t =
-      Const (@{const_name bounded_forall}, bounded_forallT T) $ lambda (Free (s, T)) t $ depth
+      Const (\<^const_name>\<open>bounded_forall\<close>, bounded_forallT T) $ lambda (Free (s, T)) t $ depth
     fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine)
     fun mk_let _ def v_opt t e = mk_let_expr (the_default def v_opt, t, e)
     val mk_test_term =
-      mk_test_term lookup mk_bounded_forall mk_safe_if mk_let @{term True} (K @{term False}) ctxt
+      mk_test_term lookup mk_bounded_forall mk_safe_if mk_let \<^term>\<open>True\<close> (K \<^term>\<open>False\<close>) ctxt
   in lambda depth (mk_test_term t) end
 
 fun mk_bounded_forall_generator_expr ctxt (t, eval_terms) =
   let
     val frees = Term.add_free_names t []
     val dummy_term =
-      @{term "Code_Evaluation.Const (STR ''Pure.dummy_pattern'') (Typerep.Typerep (STR ''dummy'') [])"}
+      \<^term>\<open>Code_Evaluation.Const (STR ''Pure.dummy_pattern'') (Typerep.Typerep (STR ''dummy'') [])\<close>
     val return =
-      @{term "Some :: term list => term list option"} $
-        (HOLogic.mk_list @{typ term} (replicate (length frees + length eval_terms) dummy_term))
-    val wrap = absdummy @{typ bool}
-      (@{term "If :: bool \<Rightarrow> term list option \<Rightarrow> term list option \<Rightarrow> term list option"} $
-        Bound 0 $ @{term "None :: term list option"} $ return)
+      \<^term>\<open>Some :: term list => term list option\<close> $
+        (HOLogic.mk_list \<^typ>\<open>term\<close> (replicate (length frees + length eval_terms) dummy_term))
+    val wrap = absdummy \<^typ>\<open>bool\<close>
+      (\<^term>\<open>If :: bool \<Rightarrow> term list option \<Rightarrow> term list option \<Rightarrow> term list option\<close> $
+        Bound 0 $ \<^term>\<open>None :: term list option\<close> $ return)
   in HOLogic.mk_comp (wrap, mk_validator_expr ctxt t) end
 
 
@@ -504,7 +504,7 @@
         (get_counterexample_batch, put_counterexample_batch,
           "Exhaustive_Generators.put_counterexample_batch")
         ctxt (SOME target) (fn proc => map (fn g => g #> (Option.map o map) proc))
-        (HOLogic.mk_list @{typ "natural => term list option"} ts') []
+        (HOLogic.mk_list \<^typ>\<open>natural \<Rightarrow> term list option\<close> ts') []
   in
     map (fn compile => fn size =>
       compile size |> (Option.map o map) Quickcheck_Common.post_process_term) compiles
@@ -518,14 +518,14 @@
   let val ts' = map (mk_validator_expr ctxt) ts in
     Code_Runtime.dynamic_value_strict
       (get_validator_batch, put_validator_batch, "Exhaustive_Generators.put_validator_batch")
-      ctxt (SOME target) (K I) (HOLogic.mk_list @{typ "natural \<Rightarrow> bool"} ts') []
+      ctxt (SOME target) (K I) (HOLogic.mk_list \<^typ>\<open>natural \<Rightarrow> bool\<close> ts') []
   end
 
 fun compile_validator_exprs ctxt ts =
   compile_validator_exprs_raw ctxt ts
   |> map (fn f => fn size => f (Code_Numeral.natural_of_integer size))
 
-fun size_matters_for thy Ts = not (forall (fn T => Sign.of_sort thy (T,  @{sort check_all})) Ts)
+fun size_matters_for thy Ts = not (forall (fn T => Sign.of_sort thy (T, \<^sort>\<open>check_all\<close>)) Ts)
 
 val test_goals =
   Quickcheck_Common.generator_test_goal_terms
@@ -535,22 +535,22 @@
 (* setup *)
 
 val setup_exhaustive_datatype_interpretation =
-  Quickcheck_Common.datatype_interpretation @{plugin quickcheck_exhaustive}
-    (@{sort exhaustive}, instantiate_exhaustive_datatype)
+  Quickcheck_Common.datatype_interpretation \<^plugin>\<open>quickcheck_exhaustive\<close>
+    (\<^sort>\<open>exhaustive\<close>, instantiate_exhaustive_datatype)
 
 val setup_bounded_forall_datatype_interpretation =
-  BNF_LFP_Compat.interpretation @{plugin quickcheck_bounded_forall} Quickcheck_Common.compat_prefs
+  BNF_LFP_Compat.interpretation \<^plugin>\<open>quickcheck_bounded_forall\<close> Quickcheck_Common.compat_prefs
     (Quickcheck_Common.ensure_sort
-       (((@{sort type}, @{sort type}), @{sort bounded_forall}),
+       (((\<^sort>\<open>type\<close>, \<^sort>\<open>type\<close>), \<^sort>\<open>bounded_forall\<close>),
        (fn thy => BNF_LFP_Compat.the_descr thy Quickcheck_Common.compat_prefs,
         instantiate_bounded_forall_datatype)))
 
-val active = Attrib.setup_config_bool @{binding quickcheck_exhaustive_active} (K true)
+val active = Attrib.setup_config_bool \<^binding>\<open>quickcheck_exhaustive_active\<close> (K true)
 
 val _ =
   Theory.setup
-   (Quickcheck_Common.datatype_interpretation @{plugin quickcheck_full_exhaustive}
-      (@{sort full_exhaustive}, instantiate_full_exhaustive_datatype)
+   (Quickcheck_Common.datatype_interpretation \<^plugin>\<open>quickcheck_full_exhaustive\<close>
+      (\<^sort>\<open>full_exhaustive\<close>, instantiate_full_exhaustive_datatype)
     #> Context.theory_map (Quickcheck.add_tester ("exhaustive", (active, test_goals)))
     #> Context.theory_map (Quickcheck.add_batch_generator ("exhaustive", compile_generator_exprs))
     #> Context.theory_map (Quickcheck.add_batch_validator ("exhaustive", compile_validator_exprs)))