src/Tools/Spec_Check/property.ML
author wenzelm
Thu, 07 Apr 2016 20:51:52 +0200
changeset 62908 d7009a515733
parent 53164 beb4ee344c22
permissions -rw-r--r--
clarified mode of ROOT.ML files;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53164
beb4ee344c22 clarified position of Spec_Check for Isabelle/ML -- it is unrelated to Isabelle/HOL;
wenzelm
parents: 52252
diff changeset
     1
(*  Title:      Tools/Spec_Check/property.ML
52248
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
Conditional properties that can track argument distribution.
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 PROPERTY =
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
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    11
type 'a pred = 'a -> bool
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    12
type 'a prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    13
val pred : 'a pred -> 'a prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    14
val pred2 : ('a * 'b) pred -> 'b -> 'a  prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    15
val implies : 'a pred * 'a prop -> 'a prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    16
val ==> : 'a pred * 'a pred -> 'a prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    17
val trivial : 'a pred -> 'a prop -> 'a prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    18
val classify : 'a pred -> string -> 'a prop -> 'a prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    19
val classify' : ('a -> string option) -> 'a prop -> 'a prop
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    20
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    21
(*Results*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    22
type result = bool option
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    23
type stats = { tags : (string * int) list, count : int }
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    24
val test : 'a prop -> 'a * stats -> result * stats
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    25
val stats : stats
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    26
val success : result pred
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    27
val failure : result pred
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    28
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    29
end
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    30
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    31
structure Property : PROPERTY =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    32
struct
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    33
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    34
type result = bool option
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    35
type stats = { tags : (string * int) list, count : int }
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    36
type 'a pred = 'a -> bool
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    37
type 'a prop = 'a * stats -> result * stats
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
fun success (SOME true) = true
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    40
  | success _ = false
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    41
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    42
fun failure (SOME false) = true
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    43
  | failure _ = false
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    44
52249
f4bf6b126cab do not handle arbitrary exceptions;
wenzelm
parents: 52248
diff changeset
    45
fun apply f x = (case try f x of NONE => SOME false | some => some)
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    46
fun pred f (x,s) = (apply f x, s)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    47
fun pred2 f z = pred (fn x => f (x, z))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    48
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    49
fun implies (cond, p) (x,s) = if cond x then p (x, s) else (NONE, s)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    50
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    51
fun ==> (p1, p2) = implies (p1, pred p2)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    52
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    53
fun wrap trans p (x,s) =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    54
  let val (result,s) = p (x,s)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    55
  in (result, trans (x, result, s)) end
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
fun classify' f =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    58
  wrap (fn (x, result, {tags, count}) =>
52252
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52249
diff changeset
    59
    { tags =
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52249
diff changeset
    60
        if is_some result then
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52249
diff changeset
    61
          (case f x of
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52249
diff changeset
    62
            NONE => tags
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52249
diff changeset
    63
          | SOME tag => AList.map_default (op =) (tag, 0) (fn c => c + 1) tags)
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52249
diff changeset
    64
        else tags,
81fcc11d8c65 minor tuning -- more linebreaks;
wenzelm
parents: 52249
diff changeset
    65
     count = count })
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    66
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    67
fun classify p tag = classify' (fn x => if p x then SOME tag else NONE)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    68
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    69
fun trivial cond = classify cond "trivial"
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    70
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    71
fun test p =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    72
  wrap (fn (_, result, {tags, count}) =>
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    73
    { tags = tags, count = if is_some result then count + 1 else count }) p
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    74
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    75
val stats = { tags = [], count = 0 }
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    76
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    77
end