src/HOL/Spec_Check/spec_check.ML
author wenzelm
Thu, 30 May 2013 21:28:54 +0200
changeset 52254 994055f7db80
parent 52253 afca6a99a361
child 52256 24f59223430d
permissions -rw-r--r--
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/Spec_Check/spec_check.ML
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     3
    Author:     Christopher League
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     4
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     5
Specification-based testing of ML programs with random values.
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     6
*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     7
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     8
signature SPEC_CHECK =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     9
sig
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    10
  val gen_target : int Config.T
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    11
  val gen_max : int Config.T
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    12
  val examples : int Config.T
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    13
  val sort_examples : bool Config.T
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    14
  val show_stats : bool Config.T
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    15
  val column_width : int Config.T
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    16
  val style : string Config.T
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    17
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    18
  type output_style = Proof.context -> string ->
52252
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
    19
    {status : string option * Property.result * (Property.stats  * string option list) -> unit,
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
    20
     finish: Property.stats * string option list -> unit}
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    21
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    22
  val register_style : (string * output_style) -> theory -> theory
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    23
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    24
  val checkGen : Proof.context ->
52252
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
    25
    'a Generator.gen * ('a -> string) option -> string * 'a Property.prop -> unit
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    26
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    27
  val check_property : Proof.context -> string -> unit
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    28
end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    29
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    30
structure Spec_Check : SPEC_CHECK =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    31
struct
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    32
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    33
(* configurations *)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    34
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    35
val gen_target = Attrib.setup_config_int @{binding spec_check_gen_target} (K 100)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    36
val gen_max = Attrib.setup_config_int @{binding spec_check_gen_max} (K 400)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    37
val examples = Attrib.setup_config_int @{binding spec_check_examples} (K 5)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    38
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    39
val sort_examples = Attrib.setup_config_bool @{binding spec_check_sort_examples} (K true)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    40
val show_stats = Attrib.setup_config_bool @{binding spec_check_show_stats} (K true)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    41
val column_width = Attrib.setup_config_int @{binding spec_check_column_width} (K 22)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    42
val style = Attrib.setup_config_string @{binding spec_check_style} (K "Perl")
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    43
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    44
open Property
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    45
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    46
type ('a, 'b) reader = 'b -> ('a * 'b) option
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    47
type 'a rep = ('a -> string) option
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    48
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    49
type output_style = Proof.context -> string ->
52252
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
    50
  {status: string option * Property.result * (Property.stats * string option list) -> unit,
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
    51
   finish: Property.stats * string option list -> unit}
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    52
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    53
fun limit ctxt gen = Generator.limit (Config.get ctxt gen_max) gen
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    54
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    55
structure Style = Theory_Data
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    56
(
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    57
  type T = output_style Symtab.table
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    58
  val empty = Symtab.empty
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    59
  val extend = I
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    60
  fun merge data : T = Symtab.merge (K true) data
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    61
)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    62
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    63
fun get_style ctxt =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    64
  let val name = Config.get ctxt style in
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    65
    (case Symtab.lookup (Style.get (Proof_Context.theory_of ctxt)) name of
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    66
      SOME style => style ctxt
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    67
    | NONE => error ("No style called " ^ quote name ^ " found"))
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    68
  end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    69
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    70
fun register_style (name, style) = Style.map (Symtab.update (name, style))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    71
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
    72
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    73
(* testing functions *)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    74
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    75
fun cpsCheck ctxt s0 shrink (next, show) (tag, prop) =
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    76
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    77
    val apply_show = case show of NONE => (fn _ => NONE) | SOME f => SOME o f
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    78
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    79
    val {status, finish} = get_style ctxt tag
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    80
    fun status' (obj, result, (stats, badobjs)) =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    81
      let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    82
        val badobjs' = if Property.failure result then obj :: badobjs else badobjs
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    83
        val _ = status (apply_show obj, result, (stats, map apply_show badobjs'))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    84
      in badobjs' end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    85
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    86
    fun try_shrink (obj, result, stats') (stats, badobjs) =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    87
      let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    88
        fun is_failure s =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    89
          let val (result, stats') = Property.test prop (s, stats)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    90
          in if Property.failure result then SOME (s, result, stats') else NONE end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    91
      in
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    92
        case get_first is_failure (shrink obj) of
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    93
          SOME (obj, result, stats') => try_shrink (obj, result, stats') (stats, badobjs)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    94
        | NONE => status' (obj, result, (stats', badobjs))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    95
      end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    96
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    97
    fun iter (NONE, (stats, badobjs)) = finish (stats, map apply_show badobjs)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    98
      | iter (SOME (obj, stream), (stats, badobjs)) =
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    99
        if #count stats >= Config.get ctxt gen_target then
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   100
          finish (stats, map apply_show badobjs)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   101
        else
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   102
          let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   103
            val (result, stats') = Property.test prop (obj, stats)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   104
            val badobjs' = if Property.failure result then
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   105
                try_shrink (obj, result, stats') (stats, badobjs)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   106
              else
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   107
                status' (obj, result, (stats', badobjs))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   108
          in iter (next stream, (stats', badobjs')) end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   109
  in
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   110
    fn stream => iter (next stream, (s0, []))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   111
  end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   112
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   113
fun check' ctxt s0 = cpsCheck ctxt s0 (fn _ => [])
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   114
fun check ctxt = check' ctxt Property.stats
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   115
fun checks ctxt = cpsCheck ctxt Property.stats
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   116
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   117
fun checkGen ctxt (gen, show) (tag, prop) =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   118
  check' ctxt {count = 0, tags = [("__GEN", 0)]}
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   119
    (limit ctxt gen, show) (tag, prop) Generator.stream
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   120
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   121
fun checkGenShrink ctxt shrink (gen, show) (tag, prop) =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   122
  cpsCheck ctxt {count=0, tags=[("__GEN", 0)]} shrink
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   123
    (limit ctxt gen, show) (tag, prop) Generator.stream
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   124
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   125
fun checkOne ctxt show (tag, prop) obj =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   126
  check ctxt (List.getItem, show) (tag, prop) [obj]
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   127
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   128
(*call the compiler and pass resulting type string to the parser*)
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   129
fun determine_type ctxt s =
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   130
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   131
    val ret = Unsynchronized.ref "return"
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   132
    val use_context : use_context = {
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   133
      tune_source = ML_Parse.fix_ints,
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   134
      name_space = ML_Env.name_space,
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   135
      str_of_pos = Position.here oo Position.line_file,
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   136
      print = fn r => ret := r,
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   137
      error = error
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   138
      }
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   139
    val _ = ctxt |> Context.proof_map
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   140
      (ML_Context.exec (fn () => Secure.use_text use_context (0, "generated code") true s))
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   141
  in
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   142
    Gen_Construction.parse_pred (!ret)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   143
  end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   144
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   145
(*call the compiler and run the test*)
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   146
fun run_test ctxt s =
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   147
  let
52252
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
   148
    val _ =
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
   149
      Context.proof_map
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   150
        (ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code")
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   151
        true s)) ctxt
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   152
  in () end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   153
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   154
(*split input into tokens*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   155
fun input_split s =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   156
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   157
    fun dot c = c = #"."
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   158
    fun space c = c = #" "
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   159
    val (head, code) = Substring.splitl (not o dot) (Substring.full s)
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   160
  in
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   161
   (String.tokens space (Substring.string head),
52252
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
   162
    Substring.string (Substring.dropl dot code))
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52248
diff changeset
   163
  end;
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   164
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   165
(*create the function from the input*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   166
fun make_fun s =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   167
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   168
    val scan_param = Scan.one (fn s => s <> ";")
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   169
    fun parameters s = Scan.repeat1 scan_param s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   170
    val p = $$ "ALL" |-- parameters
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   171
    val (split, code) = input_split s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   172
    val stop = Scan.stopper (fn _ => ";") (fn s => s = ";");
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   173
    val (params, _) = Scan.finite stop p split
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   174
  in "fn (" ^ commas params ^ ") => " ^ code end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   175
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   176
(*read input and perform the test*)
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   177
fun gen_check_property check ctxt s =
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   178
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   179
    val func = make_fun s
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   180
    val (_, ty) = determine_type ctxt func
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   181
  in run_test ctxt (check ctxt "Check" (ty, func)) end;
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   182
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   183
val check_property = gen_check_property Gen_Construction.build_check
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   184
(*val check_property_safe = gen_check_property Gen_Construction.safe_check*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   185
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   186
(*perform test for specification function*)
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   187
fun gen_check_property_f check ctxt s =
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   188
  let
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52252
diff changeset
   189
    val (name, ty) = determine_type ctxt s
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   190
  in run_test ctxt (check ctxt name (ty, s)) end;
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   191
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   192
val check_property_f = gen_check_property_f Gen_Construction.build_check
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   193
(*val check_property_safe_f_ = gen_check_property_f Gen_Construction.safe_check*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   194
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   195
end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   196
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   197
val check_property = Spec_Check.check_property;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   198