src/HOL/Tools/Quickcheck/narrowing_generators.ML
author bulwahn
Wed, 23 Mar 2011 08:50:39 +0100
changeset 42090 ef566ce50170
parent 42039 cef738d55348
child 42159 234ec7011e5d
permissions -rw-r--r--
changing Quickcheck_Narrowing's main function to enumerate the depth instead upto the depth
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:
42027
5e40c152c396 extending the test data generators to take the evaluation terms as arguments
bulwahn
parents: 42024
diff changeset
    10
    Proof.context -> term * term list -> int -> term list option * Quickcheck.report option
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    11
  val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    12
  val finite_functions : bool Config.T
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    13
  val setup: theory -> theory
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    14
end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    15
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    16
structure Narrowing_Generators : NARROWING_GENERATORS =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    17
struct
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    18
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    19
(* configurations *)
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    20
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    21
val (finite_functions, setup_finite_functions) =
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    22
  Attrib.config_bool "quickcheck_finite_functions" (K true)
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    23
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    24
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    25
fun mk_undefined T = Const(@{const_name undefined}, T)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    26
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    27
(* narrowing specific names and types *)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    28
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    29
exception FUNCTION_TYPE;
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    30
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    31
val narrowingN = "narrowing";
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    32
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    33
fun narrowingT T =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    34
  @{typ Quickcheck_Narrowing.code_int} --> Type (@{type_name Quickcheck_Narrowing.cons}, [T])
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    35
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    36
fun mk_empty T = Const (@{const_name Quickcheck_Narrowing.empty}, narrowingT T)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    37
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    38
fun mk_cons c T = Const (@{const_name Quickcheck_Narrowing.cons}, T --> narrowingT T) $ Const (c, T)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    39
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    40
fun mk_apply (T, t) (U, u) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    41
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    42
    val (_, U') = dest_funT U
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    43
  in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    44
    (U', Const (@{const_name Quickcheck_Narrowing.apply},
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    45
      narrowingT U --> narrowingT T --> narrowingT U') $ u $ t)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    46
  end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    47
  
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    48
fun mk_sum (t, u) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    49
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    50
    val T = fastype_of t
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    51
  in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    52
    Const (@{const_name Quickcheck_Narrowing.sum}, T --> T --> T) $ t $ u
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    53
  end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    54
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    55
(* creating narrowing instances *)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    56
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    57
fun mk_equations descr vs tycos narrowings (Ts, Us) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    58
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    59
    fun mk_call T =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    60
      let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    61
        val narrowing =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    62
          Const (@{const_name "Quickcheck_Narrowing.narrowing_class.narrowing"}, narrowingT T)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    63
      in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    64
        (T, narrowing)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    65
      end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    66
    fun mk_aux_call fTs (k, _) (tyco, Ts) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    67
      let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    68
        val T = Type (tyco, Ts)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    69
        val _ = if not (null fTs) then raise FUNCTION_TYPE else ()
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    70
      in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    71
        (T, nth narrowings k)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    72
      end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    73
    fun mk_consexpr simpleT (c, xs) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    74
      let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    75
        val Ts = map fst xs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    76
      in snd (fold mk_apply xs (Ts ---> simpleT, mk_cons c (Ts ---> simpleT))) end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    77
    fun mk_rhs exprs = foldr1 mk_sum exprs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    78
    val rhss =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    79
      Datatype_Aux.interpret_construction descr vs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    80
        { atyp = mk_call, dtyp = mk_aux_call }
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    81
      |> (map o apfst) Type
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    82
      |> map (fn (T, cs) => map (mk_consexpr T) cs)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    83
      |> map mk_rhs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    84
    val lhss = narrowings
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    85
    val eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhss ~~ rhss)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    86
  in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    87
    eqs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    88
  end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    89
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    90
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    91
fun instantiate_narrowing_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    92
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    93
    val _ = Datatype_Aux.message config "Creating narrowing generators ...";
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    94
    val narrowingsN = map (prefix (narrowingN ^ "_")) (names @ auxnames);
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    95
  in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    96
    thy
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    97
    |> Class.instantiation (tycos, vs, @{sort narrowing})
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    98
    |> (fold_map (fn (name, T) => Local_Theory.define
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
    99
          ((Binding.conceal (Binding.name name), NoSyn),
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   100
            (apfst Binding.conceal Attrib.empty_binding, mk_undefined (narrowingT T)))
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   101
        #> apfst fst) (narrowingsN ~~ (Ts @ Us))
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   102
      #> (fn (narrowings, lthy) =>
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   103
        let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   104
          val eqs_t = mk_equations descr vs tycos narrowings (Ts, Us)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   105
          val eqs = map (fn eq => Goal.prove lthy ["f", "i"] [] eq
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   106
            (fn _ => Skip_Proof.cheat_tac (ProofContext.theory_of lthy))) eqs_t
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   107
        in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   108
          fold (fn (name, eq) => Local_Theory.note
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   109
          ((Binding.conceal (Binding.qualify true prfx
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   110
             (Binding.qualify true name (Binding.name "simps"))),
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   111
             Code.add_default_eqn_attrib :: map (Attrib.internal o K)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   112
               [Simplifier.simp_add, Nitpick_Simps.add]), [eq]) #> snd) (narrowingsN ~~ eqs) lthy
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   113
        end))
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   114
    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   115
  end;
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   116
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   117
(* testing framework *)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   118
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   119
val target = "Haskell"
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   120
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   121
(* invocation of Haskell interpreter *)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   122
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   123
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
   124
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   125
fun exec verbose code =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   126
  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
   127
42019
a9445d87bc3e adding size as static argument in quickcheck_narrowing compilation
bulwahn
parents: 41963
diff changeset
   128
fun value ctxt (get, put, put_ml) (code, value) size =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   129
  let
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   130
    val tmp_prefix = "Quickcheck_Narrowing"
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   131
    fun run in_path = 
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   132
      let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   133
        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
   134
        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
   135
        val main_file = Path.append in_path (Path.basic "Main.hs")
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   136
        val main = "module Main where {\n\n" ^
41933
10f254a4e5b9 adapting Main file generation for Quickcheck_Narrowing
bulwahn
parents: 41932
diff changeset
   137
          "import Narrowing_Engine;\n" ^
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   138
          "import Code;\n\n" ^
42090
ef566ce50170 changing Quickcheck_Narrowing's main function to enumerate the depth instead upto the depth
bulwahn
parents: 42039
diff changeset
   139
          "main = Narrowing_Engine.depthCheck " ^ string_of_int size ^ " (Code.value ())\n\n" ^
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   140
          "}\n"
41909
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
   141
        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
   142
          (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
   143
        val _ = File.write code_file code'
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   144
        val _ = File.write narrowing_engine_file narrowing_engine
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   145
        val _ = File.write main_file main
41946
380f7f5ff126 proper File.shell_path;
wenzelm
parents: 41940
diff changeset
   146
        val executable = File.shell_path (Path.append in_path (Path.basic "isa_lsc"))
42039
cef738d55348 another attempt to exec ISABELLE_GHC robustly (cf. d8c3b26b3da4, 994d088fbfbc);
wenzelm
parents: 42028
diff changeset
   147
        val cmd = "( exec \"$ISABELLE_GHC\" -fglasgow-exts " ^
41946
380f7f5ff126 proper File.shell_path;
wenzelm
parents: 41940
diff changeset
   148
          (space_implode " " (map File.shell_path [code_file, narrowing_engine_file, main_file])) ^
42039
cef738d55348 another attempt to exec ISABELLE_GHC robustly (cf. d8c3b26b3da4, 994d088fbfbc);
wenzelm
parents: 42028
diff changeset
   149
          " -o " ^ executable ^ "; ) && " ^ executable
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   150
      in
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   151
        bash_output cmd
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   152
      end
41909
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
   153
    val result = Isabelle_System.with_tmp_dir tmp_prefix run
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   154
    val output_value = the_default "NONE"
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   155
      (try (snd o split_last o filter_out (fn s => s = "") o split_lines o fst) result)
42020
2da02764d523 translating bash output in quickcheck_narrowing to handle special characters; adding simple test cases
bulwahn
parents: 42019
diff changeset
   156
      |> translate_string (fn s => if s = "\\" then "\\\\" else s)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   157
    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
   158
      ^ " (fn () => " ^ output_value ^ ")) (ML_Context.the_generic_context ())))";
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   159
    val ctxt' = ctxt
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   160
      |> put (fn () => error ("Bad evaluation for " ^ quote put_ml))
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   161
      |> Context.proof_map (exec false ml_code);
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   162
  in get ctxt' () end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   163
42019
a9445d87bc3e adding size as static argument in quickcheck_narrowing compilation
bulwahn
parents: 41963
diff changeset
   164
fun evaluation cookie thy evaluator vs_t args size =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   165
  let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   166
    val ctxt = ProofContext.init_global thy;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   167
    val (program_code, value_name) = evaluator vs_t;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   168
    val value_code = space_implode " "
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   169
      (value_name :: "()" :: map (enclose "(" ")") args);
42019
a9445d87bc3e adding size as static argument in quickcheck_narrowing compilation
bulwahn
parents: 41963
diff changeset
   170
  in Exn.interruptible_capture (value ctxt cookie (program_code, value_code)) size end;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   171
42019
a9445d87bc3e adding size as static argument in quickcheck_narrowing compilation
bulwahn
parents: 41963
diff changeset
   172
fun dynamic_value_strict cookie thy postproc t args size =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   173
  let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   174
    fun evaluator naming program ((_, vs_ty), t) deps =
42019
a9445d87bc3e adding size as static argument in quickcheck_narrowing compilation
bulwahn
parents: 41963
diff changeset
   175
      evaluation cookie thy (Code_Target.evaluator thy target naming program deps) (vs_ty, t) args size;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   176
  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
   177
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   178
(* counterexample generator *)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   179
  
41932
e8f113ce8a94 adapting Quickcheck_Narrowing and example file to new names
bulwahn
parents: 41930
diff changeset
   180
structure Counterexample = Proof_Data
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   181
(
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   182
  type T = unit -> term list option
41936
9792a882da9c renaming tester from lazy_exhaustive to narrowing
bulwahn
parents: 41933
diff changeset
   183
  fun init _ () = error "Counterexample"
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   184
)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   185
42024
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   186
val put_counterexample = Counterexample.put
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
   187
42024
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   188
fun finitize_functions t =
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   189
  let
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   190
    val ((names, Ts), t') = apfst split_list (strip_abs t)
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   191
    fun mk_eval_ffun dT rT =
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   192
      Const (@{const_name "Quickcheck_Narrowing.eval_ffun"}, 
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   193
        Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT]) --> dT --> rT)
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   194
    fun mk_eval_cfun dT rT =
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   195
      Const (@{const_name "Quickcheck_Narrowing.eval_cfun"}, 
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   196
        Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT]) --> dT --> rT)
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   197
    fun eval_function (T as Type (@{type_name fun}, [dT, rT])) =
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   198
      let
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   199
        val (rt', rT') = eval_function rT
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   200
      in
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   201
        case dT of
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   202
          Type (@{type_name fun}, _) =>
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   203
            (fn t => absdummy (dT, rt' (mk_eval_cfun dT rT' $ incr_boundvars 1 t $ Bound 0)),
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   204
            Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT']))
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   205
        | _ => (fn t => absdummy (dT, rt' (mk_eval_ffun dT rT' $ incr_boundvars 1 t $ Bound 0)),
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   206
            Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT']))
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   207
      end
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   208
      | eval_function T = (I, T)
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   209
    val (tt, Ts') = split_list (map eval_function Ts)
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   210
    val t'' = subst_bounds (map2 (fn f => fn x => f x) (rev tt) (map_index (Bound o fst) Ts), t')
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   211
  in
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   212
    list_abs (names ~~ Ts', t'')
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   213
  end
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
   214
42027
5e40c152c396 extending the test data generators to take the evaluation terms as arguments
bulwahn
parents: 42024
diff changeset
   215
fun compile_generator_expr ctxt (t, eval_terms) size =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   216
  let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   217
    val thy = ProofContext.theory_of ctxt
42028
bd6515e113b2 passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
bulwahn
parents: 42027
diff changeset
   218
    val t' = list_abs_free (Term.add_frees t [], t)
bd6515e113b2 passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
bulwahn
parents: 42027
diff changeset
   219
    val t'' = if Config.get ctxt finite_functions then finitize_functions t' else t'
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   220
    fun ensure_testable t =
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   221
      Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
   222
    val result = dynamic_value_strict
41932
e8f113ce8a94 adapting Quickcheck_Narrowing and example file to new names
bulwahn
parents: 41930
diff changeset
   223
      (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample")
42028
bd6515e113b2 passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
bulwahn
parents: 42027
diff changeset
   224
      thy (Option.map o map) (ensure_testable t'') [] size
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   225
  in
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
   226
    (result, NONE)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   227
  end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   228
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   229
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   230
val setup =
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   231
  Datatype.interpretation
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   232
    (Quickcheck_Common.ensure_sort_datatype (@{sort narrowing}, instantiate_narrowing_datatype))
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
   233
  #> setup_finite_functions
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   234
  #> Context.theory_map
41936
9792a882da9c renaming tester from lazy_exhaustive to narrowing
bulwahn
parents: 41933
diff changeset
   235
    (Quickcheck.add_generator ("narrowing", compile_generator_expr))
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   236
    
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   237
end;