src/HOL/Tools/Quickcheck/narrowing_generators.ML
author bulwahn
Mon, 19 Sep 2011 16:18:33 +0200
changeset 45864 df36896aae0f
parent 45797 de3ed037c9a5
child 45904 632a033616e9
permissions -rw-r--r--
removing superfluous definition in the quickcheck narrowing invocation as the code generator now generates valid Haskell code with necessary type annotations without a separate definition
     1 (*  Title:      HOL/Tools/Quickcheck/narrowing_generators.ML
     2     Author:     Lukas Bulwahn, TU Muenchen
     3 
     4 Narrowing-based counterexample generation.
     5 *)
     6 
     7 signature NARROWING_GENERATORS =
     8 sig
     9   val allow_existentials : bool Config.T
    10   val finite_functions : bool Config.T
    11   val overlord : bool Config.T
    12   val active : bool Config.T
    13   val test_term: Proof.context -> bool * bool -> term * term list -> Quickcheck.result
    14   datatype counterexample = Universal_Counterexample of (term * counterexample)
    15     | Existential_Counterexample of (term * counterexample) list
    16     | Empty_Assignment
    17   val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context
    18   val put_existential_counterexample : (unit -> counterexample option) -> Proof.context -> Proof.context
    19   val setup: theory -> theory
    20 end;
    21 
    22 structure Narrowing_Generators : NARROWING_GENERATORS =
    23 struct
    24 
    25 (* configurations *)
    26 
    27 val allow_existentials = Attrib.setup_config_bool @{binding quickcheck_allow_existentials} (K true)
    28 val finite_functions = Attrib.setup_config_bool @{binding quickcheck_finite_functions} (K true)
    29 val overlord = Attrib.setup_config_bool @{binding quickcheck_narrowing_overlord} (K false)
    30 
    31 (* partial_term_of instances *)
    32 
    33 fun mk_partial_term_of (x, T) =
    34   Const (@{const_name Quickcheck_Narrowing.partial_term_of_class.partial_term_of},
    35     Term.itselfT T --> @{typ narrowing_term} --> @{typ Code_Evaluation.term})
    36       $ Const ("TYPE", Term.itselfT T) $ x
    37 
    38 (** formal definition **)
    39 
    40 fun add_partial_term_of tyco raw_vs thy =
    41   let
    42     val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
    43     val ty = Type (tyco, map TFree vs);
    44     val lhs = Const (@{const_name partial_term_of},
    45         Term.itselfT ty --> @{typ narrowing_term} --> @{typ Code_Evaluation.term})
    46       $ Free ("x", Term.itselfT ty) $ Free ("t", @{typ narrowing_term});
    47     val rhs = @{term "undefined :: Code_Evaluation.term"};
    48     val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
    49     fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
    50       o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
    51   in
    52     thy
    53     |> Class.instantiation ([tyco], vs, @{sort partial_term_of})
    54     |> `(fn lthy => Syntax.check_term lthy eq)
    55     |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
    56     |> snd
    57     |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
    58   end;
    59 
    60 fun ensure_partial_term_of (tyco, (raw_vs, _)) thy =
    61   let
    62     val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort partial_term_of})
    63       andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
    64   in if need_inst then add_partial_term_of tyco raw_vs thy else thy end;
    65 
    66 
    67 (** code equations for datatypes **)
    68 
    69 fun mk_partial_term_of_eq thy ty (i, (c, (_, tys))) =
    70   let
    71     val frees = map Free (Name.invent_names Name.context "a" (map (K @{typ narrowing_term}) tys))
    72     val narrowing_term = @{term "Quickcheck_Narrowing.Ctr"} $ HOLogic.mk_number @{typ code_int} i
    73       $ (HOLogic.mk_list @{typ narrowing_term} (rev frees))
    74     val rhs = fold (fn u => fn t => @{term "Code_Evaluation.App"} $ t $ u)
    75         (map mk_partial_term_of (frees ~~ tys))
    76         (@{term "Code_Evaluation.Const"} $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> ty))
    77     val insts =
    78       map (SOME o Thm.cterm_of thy o map_types Logic.unvarifyT_global o Logic.varify_global)
    79         [Free ("ty", Term.itselfT ty), narrowing_term, rhs]
    80     val cty = Thm.ctyp_of thy ty;
    81   in
    82     @{thm partial_term_of_anything}
    83     |> Drule.instantiate' [SOME cty] insts
    84     |> Thm.varifyT_global
    85   end
    86 
    87 fun add_partial_term_of_code tyco raw_vs raw_cs thy =
    88   let
    89     val algebra = Sign.classes_of thy;
    90     val vs = map (fn (v, sort) =>
    91       (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
    92     val ty = Type (tyco, map TFree vs);
    93     val cs = (map o apsnd o apsnd o map o map_atyps)
    94       (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
    95     val const = AxClass.param_of_inst thy (@{const_name partial_term_of}, tyco);
    96     val var_insts = map (SOME o Thm.cterm_of thy o map_types Logic.unvarifyT_global o Logic.varify_global)
    97         [Free ("ty", Term.itselfT ty), @{term "Quickcheck_Narrowing.Var p tt"},
    98           @{term "Code_Evaluation.Free (STR ''_'')"} $ HOLogic.mk_typerep ty]
    99     val var_eq =
   100       @{thm partial_term_of_anything}
   101       |> Drule.instantiate' [SOME (Thm.ctyp_of thy ty)] var_insts
   102       |> Thm.varifyT_global
   103     val eqs = var_eq :: map_index (mk_partial_term_of_eq thy ty) cs;
   104  in
   105     thy
   106     |> Code.del_eqns const
   107     |> fold Code.add_eqn eqs
   108   end;
   109 
   110 fun ensure_partial_term_of_code (tyco, (raw_vs, cs)) thy =
   111   let
   112     val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort partial_term_of};
   113   in if has_inst then add_partial_term_of_code tyco raw_vs cs thy else thy end;
   114 
   115 
   116 (* narrowing generators *)
   117 
   118 (** narrowing specific names and types **)
   119 
   120 exception FUNCTION_TYPE;
   121 
   122 val narrowingN = "narrowing";
   123 
   124 fun narrowingT T =
   125   @{typ Quickcheck_Narrowing.code_int} --> Type (@{type_name Quickcheck_Narrowing.cons}, [T])
   126 
   127 fun mk_empty T = Const (@{const_name Quickcheck_Narrowing.empty}, narrowingT T)
   128 
   129 fun mk_cons c T = Const (@{const_name Quickcheck_Narrowing.cons}, T --> narrowingT T) $ Const (c, T)
   130 
   131 fun mk_apply (T, t) (U, u) =
   132   let
   133     val (_, U') = dest_funT U
   134   in
   135     (U', Const (@{const_name Quickcheck_Narrowing.apply},
   136       narrowingT U --> narrowingT T --> narrowingT U') $ u $ t)
   137   end
   138   
   139 fun mk_sum (t, u) =
   140   let
   141     val T = fastype_of t
   142   in
   143     Const (@{const_name Quickcheck_Narrowing.sum}, T --> T --> T) $ t $ u
   144   end
   145 
   146 (** deriving narrowing instances **)
   147 
   148 fun mk_equations descr vs tycos narrowings (Ts, Us) =
   149   let
   150     fun mk_call T =
   151       (T, Const (@{const_name "Quickcheck_Narrowing.narrowing_class.narrowing"}, narrowingT T))
   152     fun mk_aux_call fTs (k, _) (tyco, Ts) =
   153       let
   154         val T = Type (tyco, Ts)
   155         val _ = if not (null fTs) then raise FUNCTION_TYPE else ()
   156       in
   157         (T, nth narrowings k)
   158       end
   159     fun mk_consexpr simpleT (c, xs) =
   160       let
   161         val Ts = map fst xs
   162       in snd (fold mk_apply xs (Ts ---> simpleT, mk_cons c (Ts ---> simpleT))) end
   163     fun mk_rhs exprs = foldr1 mk_sum exprs
   164     val rhss =
   165       Datatype_Aux.interpret_construction descr vs
   166         { atyp = mk_call, dtyp = mk_aux_call }
   167       |> (map o apfst) Type
   168       |> map (fn (T, cs) => map (mk_consexpr T) cs)
   169       |> map mk_rhs
   170     val lhss = narrowings
   171     val eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhss ~~ rhss)
   172   in
   173     eqs
   174   end
   175     
   176 fun contains_recursive_type_under_function_types xs =
   177   exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT =>
   178     (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype.DtRec _) => true | _ => false)))) xs
   179 
   180 fun instantiate_narrowing_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
   181   let
   182     val _ = Datatype_Aux.message config "Creating narrowing generators ...";
   183     val narrowingsN = map (prefix (narrowingN ^ "_")) (names @ auxnames);
   184   in
   185     if not (contains_recursive_type_under_function_types descr) then
   186       thy
   187       |> Class.instantiation (tycos, vs, @{sort narrowing})
   188       |> Quickcheck_Common.define_functions
   189         (fn narrowings => mk_equations descr vs tycos narrowings (Ts, Us), NONE)
   190         prfx [] narrowingsN (map narrowingT (Ts @ Us))
   191       |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
   192     else
   193       thy
   194   end;
   195 
   196 (* testing framework *)
   197 
   198 val target = "Haskell_Quickcheck"
   199 
   200 (** invocation of Haskell interpreter **)
   201 
   202 val narrowing_engine =
   203   Context.>>> (Context.map_theory_result
   204     (Thy_Load.use_file (Path.explode "Tools/Quickcheck/Narrowing_Engine.hs")))
   205 
   206 val pnf_narrowing_engine =
   207   Context.>>> (Context.map_theory_result
   208     (Thy_Load.use_file (Path.explode "Tools/Quickcheck/PNF_Narrowing_Engine.hs")))
   209 
   210 fun exec verbose code =
   211   ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code)
   212 
   213 fun with_overlord_dir name f =
   214   let
   215     val path =
   216       Path.append (Path.explode "$ISABELLE_HOME_USER") (Path.basic (name ^ serial_string ()))
   217     val _ = Isabelle_System.mkdirs path;
   218   in Exn.release (Exn.capture f path) end;
   219 
   220 fun elapsed_time description e =
   221   let val ({elapsed, ...}, result) = Timing.timing e ()
   222   in (result, (description, Time.toMilliseconds elapsed)) end
   223   
   224 fun value (contains_existentials, opts) ctxt cookie (code, value_name) =
   225   let
   226     val (limit_time, is_interactive, timeout, quiet, size) = opts
   227     val (get, put, put_ml) = cookie
   228     fun message s = if quiet then () else Output.urgent_message s
   229     val current_size = Unsynchronized.ref 0
   230     val current_result = Unsynchronized.ref Quickcheck.empty_result 
   231     fun excipit () =
   232       "Quickcheck ran out of time while testing at size " ^ string_of_int (!current_size)
   233     val tmp_prefix = "Quickcheck_Narrowing"
   234     val with_tmp_dir =
   235       if Config.get ctxt overlord then with_overlord_dir else Isabelle_System.with_tmp_dir 
   236     fun run in_path = 
   237       let
   238         val code_file = Path.append in_path (Path.basic "Generated_Code.hs")
   239         val narrowing_engine_file = Path.append in_path (Path.basic "Narrowing_Engine.hs")
   240         val main_file = Path.append in_path (Path.basic "Main.hs")
   241         val main = "module Main where {\n\n" ^
   242           "import System;\n" ^
   243           "import Narrowing_Engine;\n" ^
   244           "import Generated_Code;\n\n" ^
   245           "main = getArgs >>= \\[size] -> Narrowing_Engine.depthCheck (read size) (Generated_Code.value ())\n\n" ^
   246           "}\n"
   247         val code' = prefix "module Generated_Code where {\n\ndata Typerep = Typerep String [Typerep];\n"
   248           (unprefix "module Generated_Code where {" code)
   249         val _ = File.write code_file code'
   250         val _ = File.write narrowing_engine_file
   251           (if contains_existentials then pnf_narrowing_engine else narrowing_engine)
   252         val _ = File.write main_file main
   253         val executable = File.shell_path (Path.append in_path (Path.basic "isabelle_quickcheck_narrowing"))
   254         val cmd = "exec \"$ISABELLE_GHC\" " ^ Code_Haskell.language_params ^ " " ^
   255           (space_implode " " (map File.shell_path [code_file, narrowing_engine_file, main_file])) ^
   256           " -o " ^ executable ^ ";"
   257         val (result, compilation_time) =
   258           elapsed_time "Haskell compilation" (fn () => Isabelle_System.bash cmd) 
   259         val _ = Quickcheck.add_timing compilation_time current_result
   260         val _ = if Isabelle_System.bash cmd <> 0 then error "Compilation with GHC failed" else ()
   261         fun with_size k =
   262           if k > size then
   263             (NONE, !current_result)
   264           else
   265             let
   266               val _ = message ("[Quickcheck-Narrowing] Test data size: " ^ string_of_int k)
   267               val _ = current_size := k
   268               val ((response, _), timing) = elapsed_time ("execution of size " ^ string_of_int k)
   269                 (fn () => Isabelle_System.bash_output (executable ^ " " ^ string_of_int k))
   270               val _ = Quickcheck.add_timing timing current_result
   271             in
   272               if response = "NONE\n" then
   273                 with_size (k + 1)
   274               else
   275                 let
   276                   val output_value = the_default "NONE"
   277                     (try (snd o split_last o filter_out (fn s => s = "") o split_lines) response)
   278                     |> translate_string (fn s => if s = "\\" then "\\\\" else s)
   279                   val ml_code = "\nval _ = Context.set_thread_data (SOME (Context.map_proof (" ^ put_ml
   280                     ^ " (fn () => " ^ output_value ^ ")) (ML_Context.the_generic_context ())))";
   281                   val ctxt' = ctxt
   282                     |> put (fn () => error ("Bad evaluation for " ^ quote put_ml))
   283                     |> Context.proof_map (exec false ml_code);
   284                 in (get ctxt' (), !current_result) end
   285             end 
   286       in with_size 0 end
   287   in
   288     (*Quickcheck.limit timeout (limit_time, is_interactive) 
   289       (fn () =>*) with_tmp_dir tmp_prefix run
   290       (*(fn () => (message (excipit ()); (NONE, !current_result))) ()*)
   291   end;
   292 
   293 fun dynamic_value_strict opts cookie thy postproc t =
   294   let
   295     val ctxt = Proof_Context.init_global thy
   296     fun evaluator naming program ((_, vs_ty), t) deps = Exn.interruptible_capture (value opts ctxt cookie)
   297       (Code_Target.evaluator thy target naming program deps (vs_ty, t));    
   298   in Exn.release (Code_Thingol.dynamic_value thy (Exn.map_result o postproc) evaluator t) end;
   299 
   300 (** counterexample generator **)
   301   
   302 structure Counterexample = Proof_Data
   303 (
   304   type T = unit -> term list option
   305   fun init _ () = error "Counterexample"
   306 )
   307 
   308 datatype counterexample = Universal_Counterexample of (term * counterexample)
   309   | Existential_Counterexample of (term * counterexample) list
   310   | Empty_Assignment
   311   
   312 fun map_counterexample f Empty_Assignment = Empty_Assignment
   313   | map_counterexample f (Universal_Counterexample (t, c)) =
   314       Universal_Counterexample (f t, map_counterexample f c)
   315   | map_counterexample f (Existential_Counterexample cs) =
   316       Existential_Counterexample (map (fn (t, c) => (f t, map_counterexample f c)) cs)
   317 
   318 structure Existential_Counterexample = Proof_Data
   319 (
   320   type T = unit -> counterexample option
   321   fun init _ () = error "Counterexample"
   322 )
   323 
   324 val put_existential_counterexample = Existential_Counterexample.put
   325 
   326 val put_counterexample = Counterexample.put
   327 
   328 fun finitize_functions (xTs, t) =
   329   let
   330     val (names, boundTs) = split_list xTs
   331     fun mk_eval_ffun dT rT =
   332       Const (@{const_name "Quickcheck_Narrowing.eval_ffun"}, 
   333         Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT]) --> dT --> rT)
   334     fun mk_eval_cfun dT rT =
   335       Const (@{const_name "Quickcheck_Narrowing.eval_cfun"}, 
   336         Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT]) --> dT --> rT)
   337     fun eval_function (T as Type (@{type_name fun}, [dT, rT])) =
   338       let
   339         val (rt', rT') = eval_function rT
   340       in
   341         case dT of
   342           Type (@{type_name fun}, _) =>
   343             (fn t => absdummy dT (rt' (mk_eval_cfun dT rT' $ incr_boundvars 1 t $ Bound 0)),
   344               Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT']))
   345         | _ =>
   346             (fn t => absdummy dT (rt' (mk_eval_ffun dT rT' $ incr_boundvars 1 t $ Bound 0)),
   347               Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT']))
   348       end
   349       | eval_function T = (I, T)
   350     val (tt, boundTs') = split_list (map eval_function boundTs)
   351     val t' = subst_bounds (map2 (fn f => fn x => f x) (rev tt) (map_index (Bound o fst) boundTs), t)
   352   in
   353     (names ~~ boundTs', t')
   354   end
   355 
   356 (** tester **)
   357 
   358 val rewrs =
   359     map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of)
   360       (@{thms all_simps} @ @{thms ex_simps})
   361     @ map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of)
   362         [@{thm iff_conv_conj_imp}, @{thm not_ex}, @{thm not_all}]
   363 
   364 fun make_pnf_term thy t = Pattern.rewrite_term thy rewrs [] t
   365 
   366 fun strip_quantifiers (Const (@{const_name Ex}, _) $ Abs (x, T, t)) =
   367     apfst (cons (@{const_name Ex}, (x, T))) (strip_quantifiers t)
   368   | strip_quantifiers (Const (@{const_name All}, _) $ Abs (x, T, t)) =
   369     apfst (cons (@{const_name All}, (x, T))) (strip_quantifiers t)
   370   | strip_quantifiers t = ([], t)
   371 
   372 fun contains_existentials t = exists (fn (Q, _) => Q = @{const_name Ex}) (fst (strip_quantifiers t))
   373 
   374 fun mk_property qs t =
   375   let
   376     fun enclose (@{const_name Ex}, (x, T)) t =
   377         Const (@{const_name Quickcheck_Narrowing.exists}, (T --> @{typ property}) --> @{typ property})
   378           $ Abs (x, T, t)
   379       | enclose (@{const_name All}, (x, T)) t =
   380         Const (@{const_name Quickcheck_Narrowing.all}, (T --> @{typ property}) --> @{typ property})
   381           $ Abs (x, T, t)
   382   in
   383     fold_rev enclose qs (@{term Quickcheck_Narrowing.Property} $ t)
   384   end
   385 
   386 fun mk_case_term ctxt p ((@{const_name Ex}, (x, T)) :: qs') (Existential_Counterexample cs) =
   387     Datatype.make_case ctxt Datatype_Case.Quiet [] (Free (x, T)) (map (fn (t, c) =>
   388       (t, mk_case_term ctxt (p - 1) qs' c)) cs)
   389   | mk_case_term ctxt p ((@{const_name All}, (x, T)) :: qs') (Universal_Counterexample (t, c)) =
   390     if p = 0 then t else mk_case_term ctxt (p - 1) qs' c
   391 
   392 fun mk_terms ctxt qs result =
   393   let
   394     val
   395       ps = filter (fn (_, (@{const_name All}, _)) => true | _ => false) (map_index I qs)
   396     in
   397       map (fn (p, (_, (x, T))) => (x, mk_case_term ctxt p qs result)) ps
   398     end
   399   
   400 fun test_term ctxt (limit_time, is_interactive) (t, eval_terms) =
   401   let
   402     fun dest_result (Quickcheck.Result r) = r 
   403     val opts =
   404       (limit_time, is_interactive, seconds (Config.get ctxt Quickcheck.timeout),
   405         Config.get ctxt Quickcheck.quiet, Config.get ctxt Quickcheck.size)
   406     val thy = Proof_Context.theory_of ctxt
   407     val t' = fold_rev (fn (x, T) => fn t => HOLogic.mk_all (x, T, t)) (Term.add_frees t []) t
   408     val pnf_t = make_pnf_term thy t'
   409   in
   410     if Config.get ctxt allow_existentials andalso contains_existentials pnf_t then
   411       let
   412         fun wrap f (qs, t) =
   413           let val (qs1, qs2) = split_list qs in
   414           apfst (map2 pair qs1) (f (qs2, t)) end
   415         val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I
   416         val (qs, prop_t) = finitize (strip_quantifiers pnf_t)
   417         val (counterexample, result) = dynamic_value_strict (true, opts)
   418           (Existential_Counterexample.get, Existential_Counterexample.put,
   419             "Narrowing_Generators.put_existential_counterexample")
   420           thy (apfst o Option.map o map_counterexample) (mk_property qs prop_t)
   421       in
   422         Quickcheck.Result
   423          {counterexample = Option.map (mk_terms ctxt qs) counterexample,
   424           evaluation_terms = Option.map (K []) counterexample,
   425           timings = #timings (dest_result result), reports = #reports (dest_result result)}
   426       end
   427     else
   428       let
   429         val t' = fold_rev absfree (Term.add_frees t []) t
   430         fun wrap f t = list_abs (f (strip_abs t))
   431         val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I
   432         fun ensure_testable t =
   433           Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t
   434         val (counterexample, result) = dynamic_value_strict (false, opts)
   435           (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample")
   436           thy (apfst o Option.map o map) (ensure_testable (finitize t'))
   437       in
   438         Quickcheck.Result
   439          {counterexample = Option.map ((curry (op ~~)) (Term.add_free_names t [])) counterexample,
   440           evaluation_terms = Option.map (K []) counterexample,
   441           timings = #timings (dest_result result), reports = #reports (dest_result result)}
   442       end
   443   end;
   444 
   445 fun test_goals ctxt (limit_time, is_interactive) insts goals =
   446   if (not (getenv "ISABELLE_GHC" = "")) then
   447     let
   448       val correct_inst_goals = Quickcheck.instantiate_goals ctxt insts goals
   449     in
   450       Quickcheck.collect_results (test_term ctxt (limit_time, is_interactive)) (maps (map snd) correct_inst_goals) []
   451     end
   452   else
   453     (if Config.get ctxt Quickcheck.quiet then () else Output.urgent_message
   454       ("Environment variable ISABELLE_GHC is not set. To use narrowing-based quickcheck, please set "
   455         ^ "this variable to your GHC Haskell compiler in your settings file. "
   456         ^ "To deactivate narrowing-based quickcheck, set quickcheck_narrowing_active to false.");
   457       [Quickcheck.empty_result])
   458 
   459 (* setup *)
   460 
   461 val active = Attrib.setup_config_bool @{binding quickcheck_narrowing_active} (K true);
   462 
   463 val setup =
   464   Code.datatype_interpretation ensure_partial_term_of
   465   #> Code.datatype_interpretation ensure_partial_term_of_code
   466   #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype
   467     (((@{sort typerep}, @{sort term_of}), @{sort narrowing}), instantiate_narrowing_datatype))
   468   #> Context.theory_map (Quickcheck.add_tester ("narrowing", (active, test_goals)))
   469     
   470 end;