author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 76183 | 8089593a364a |
permissions | -rw-r--r-- |
76183 | 1 |
(* Title: HOL/Tools/Mirabelle/mirabelle_quickcheck.ML |
73847
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
2 |
Author: Jasmin Blanchette, TU Munich |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
3 |
Author: Sascha Boehme, TU Munich |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
4 |
Author: Martin Desharnais, UniBw Munich |
73691
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
wenzelm
parents:
67405
diff
changeset
|
5 |
|
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
wenzelm
parents:
67405
diff
changeset
|
6 |
Mirabelle action: "quickcheck". |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
7 |
*) |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
8 |
|
73847
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
9 |
structure Mirabelle_Quickcheck: MIRABELLE_ACTION = |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
10 |
struct |
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
11 |
|
73847
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
12 |
fun make_action ({arguments, timeout, ...} : Mirabelle.action_context) = |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
13 |
let |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
14 |
val has_valid_key = member (op =) ["iterations", "size", "generator"] o fst |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
15 |
val quickcheck = |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
16 |
Quickcheck.quickcheck (map (apsnd single) (filter has_valid_key arguments)) 1 |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
17 |
|
75003
f21e7e6172a0
renamed run_action to run in Mirabelle.action record
desharna
parents:
74948
diff
changeset
|
18 |
fun run ({pre, ...} : Mirabelle.command) = |
73847
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
19 |
(case Timeout.apply timeout quickcheck pre of |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
20 |
NONE => "no counterexample" |
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
21 |
| SOME _ => "counterexample found") |
75003
f21e7e6172a0
renamed run_action to run in Mirabelle.action record
desharna
parents:
74948
diff
changeset
|
22 |
in ("", {run = run, finalize = K ""}) end |
73847
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
23 |
|
58f6b41efe88
refactored Mirabelle to produce output in real time
desharna
parents:
73691
diff
changeset
|
24 |
val () = Mirabelle.register_action "quickcheck" make_action |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
25 |
|
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
26 |
end |