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