src/HOL/Tools/Quickcheck/narrowing_generators.ML
author wenzelm
Sun, 13 Mar 2011 19:16:19 +0100
changeset 41952 c7297638599b
parent 41946 380f7f5ff126
child 41953 994d088fbfbc
permissions -rw-r--r--
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection; determine swipl_version at runtime;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
     1
(*  Title:      HOL/Tools/Quickcheck/narrowing_generators.ML
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn, TU Muenchen
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     3
41938
645cca858c69 tuned headers;
wenzelm
parents: 41936
diff changeset
     4
Narrowing-based counterexample generation.
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     5
*)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     6
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
     7
signature NARROWING_GENERATORS =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     8
sig
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     9
  val compile_generator_expr:
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    10
    Proof.context -> term -> int -> term list option * Quickcheck.report option
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    11
  val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    12
  val setup: theory -> theory
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    13
end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    14
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    15
structure Narrowing_Generators : NARROWING_GENERATORS =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    16
struct
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    17
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    18
val target = "Haskell"
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    19
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    20
(* invocation of Haskell interpreter *)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    21
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    22
val narrowing_engine = File.read (Path.explode "~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs")
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    23
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    24
fun exec verbose code =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    25
  ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    26
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    27
fun value ctxt (get, put, put_ml) (code, value) =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    28
  let
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    29
    val tmp_prefix = "Quickcheck_Narrowing"
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    30
    fun run in_path = 
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    31
      let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    32
        val code_file = Path.append in_path (Path.basic "Code.hs")
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    33
        val narrowing_engine_file = Path.append in_path (Path.basic "Narrowing_Engine.hs")
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    34
        val main_file = Path.append in_path (Path.basic "Main.hs")
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    35
        val main = "module Main where {\n\n" ^
41933
10f254a4e5b9 adapting Main file generation for Quickcheck_Narrowing
bulwahn
parents: 41932
diff changeset
    36
          "import Narrowing_Engine;\n" ^
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    37
          "import Code;\n\n" ^
41933
10f254a4e5b9 adapting Main file generation for Quickcheck_Narrowing
bulwahn
parents: 41932
diff changeset
    38
          "main = Narrowing_Engine.smallCheck 7 (Code.value ())\n\n" ^
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    39
          "}\n"
41909
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
    40
        val code' = prefix "module Code where {\n\ndata Typerep = Typerep String [Typerep];\n"
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
    41
          (unprefix "module Code where {" code)
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
    42
        val _ = File.write code_file code'
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    43
        val _ = File.write narrowing_engine_file narrowing_engine
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    44
        val _ = File.write main_file main
41946
380f7f5ff126 proper File.shell_path;
wenzelm
parents: 41940
diff changeset
    45
        val executable = File.shell_path (Path.append in_path (Path.basic "isa_lsc"))
41952
c7297638599b cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
wenzelm
parents: 41946
diff changeset
    46
        val cmd = "\"$ISABELLE_GHC\" -fglasgow-exts " ^
41946
380f7f5ff126 proper File.shell_path;
wenzelm
parents: 41940
diff changeset
    47
          (space_implode " " (map File.shell_path [code_file, narrowing_engine_file, main_file])) ^
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    48
          " -o " ^ executable ^ " && " ^ executable
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    49
      in
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    50
        bash_output cmd
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    51
      end
41909
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
    52
    val result = Isabelle_System.with_tmp_dir tmp_prefix run
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    53
    val output_value = the_default "NONE"
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    54
      (try (snd o split_last o filter_out (fn s => s = "") o split_lines o fst) result)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    55
    val ml_code = "\nval _ = Context.set_thread_data (SOME (Context.map_proof (" ^ put_ml
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    56
      ^ " (fn () => " ^ output_value ^ ")) (ML_Context.the_generic_context ())))";
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    57
    val ctxt' = ctxt
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    58
      |> put (fn () => error ("Bad evaluation for " ^ quote put_ml))
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    59
      |> Context.proof_map (exec false ml_code);
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    60
  in get ctxt' () end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    61
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    62
fun evaluation cookie thy evaluator vs_t args =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    63
  let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    64
    val ctxt = ProofContext.init_global thy;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    65
    val (program_code, value_name) = evaluator vs_t;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    66
    val value_code = space_implode " "
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    67
      (value_name :: "()" :: map (enclose "(" ")") args);
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    68
  in Exn.interruptible_capture (value ctxt cookie) (program_code, value_code) end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    69
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    70
fun dynamic_value_strict cookie thy postproc t args =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    71
  let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    72
    fun evaluator naming program ((_, vs_ty), t) deps =
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    73
      evaluation cookie thy (Code_Target.evaluator thy target naming program deps) (vs_ty, t) args;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    74
  in Exn.release (Code_Thingol.dynamic_value thy (Exn.map_result o postproc) evaluator t) end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    75
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    76
(* counterexample generator *)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    77
  
41932
e8f113ce8a94 adapting Quickcheck_Narrowing and example file to new names
bulwahn
parents: 41930
diff changeset
    78
structure Counterexample = Proof_Data
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    79
(
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    80
  type T = unit -> term list option
41936
9792a882da9c renaming tester from lazy_exhaustive to narrowing
bulwahn
parents: 41933
diff changeset
    81
  fun init _ () = error "Counterexample"
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    82
)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    83
41932
e8f113ce8a94 adapting Quickcheck_Narrowing and example file to new names
bulwahn
parents: 41930
diff changeset
    84
val put_counterexample =  Counterexample.put
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    85
  
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    86
fun compile_generator_expr ctxt t size =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    87
  let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    88
    val thy = ProofContext.theory_of ctxt
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    89
    fun ensure_testable t =
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    90
      Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    91
    val t = dynamic_value_strict
41932
e8f113ce8a94 adapting Quickcheck_Narrowing and example file to new names
bulwahn
parents: 41930
diff changeset
    92
      (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample")
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    93
      thy (Option.map o map) (ensure_testable t) []
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    94
  in
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    95
    (t, NONE)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    96
  end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    97
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    98
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    99
val setup =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   100
  Context.theory_map
41936
9792a882da9c renaming tester from lazy_exhaustive to narrowing
bulwahn
parents: 41933
diff changeset
   101
    (Quickcheck.add_generator ("narrowing", compile_generator_expr))
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   102
    
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   103
end;