src/Tools/auto_tools.ML
author bulwahn
Fri, 18 Mar 2011 18:19:42 +0100
changeset 42025 cb5b1e85b32e
parent 40931 061b8257ab9f
permissions -rw-r--r--
adding eval option to quickcheck
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39324
05452dd66b2b finished renaming "Auto_Counterexample" to "Auto_Tools"
blanchet
parents: 39323
diff changeset
     1
(*  Title:      Tools/auto_tools.ML
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
     3
39324
05452dd66b2b finished renaming "Auto_Counterexample" to "Auto_Tools"
blanchet
parents: 39323
diff changeset
     4
Manager for tools that should be run automatically on newly entered conjectures.
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
     5
*)
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
     6
39324
05452dd66b2b finished renaming "Auto_Counterexample" to "Auto_Tools"
blanchet
parents: 39323
diff changeset
     7
signature AUTO_TOOLS =
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
     8
sig
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
     9
  val time_limit: real Unsynchronized.ref
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    10
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    11
  val register_tool :
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    12
    bool Unsynchronized.ref * (Proof.state -> bool * Proof.state) -> theory
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    13
    -> theory
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    14
end;
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    15
39324
05452dd66b2b finished renaming "Auto_Counterexample" to "Auto_Tools"
blanchet
parents: 39323
diff changeset
    16
structure Auto_Tools : AUTO_TOOLS =
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    17
struct
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    18
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    19
(* preferences *)
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    20
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    21
val time_limit = Unsynchronized.ref 4.0
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    22
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    23
val auto_tools_time_limitN = "auto-tools-time-limit"
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    24
val _ =
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    25
  ProofGeneralPgip.add_preference Preferences.category_tracing
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    26
    (Preferences.real_pref time_limit
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    27
      auto_tools_time_limitN "Time limit for automatic tools (in seconds).")
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    28
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    29
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    30
(* configuration *)
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    31
33600
16a263d2b1c9 adapted Theory_Data;
wenzelm
parents: 33561
diff changeset
    32
structure Data = Theory_Data
16a263d2b1c9 adapted Theory_Data;
wenzelm
parents: 33561
diff changeset
    33
(
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    34
  type T = (bool Unsynchronized.ref * (Proof.state -> bool * Proof.state)) list
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    35
  val empty = []
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    36
  val extend = I
33600
16a263d2b1c9 adapted Theory_Data;
wenzelm
parents: 33561
diff changeset
    37
  fun merge data : T = AList.merge (op =) (K true) data
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    38
)
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    39
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    40
val register_tool = Data.map o AList.update (op =)
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    41
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    42
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    43
(* automatic testing *)
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    44
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    45
fun run_tools tools state =
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    46
  tools |> map_filter (fn (auto, tool) => if !auto then SOME tool else NONE)
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    47
        |> Par_List.get_some (fn tool =>
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    48
                                 case try tool state of
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    49
                                   SOME (true, state) => SOME state
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    50
                                 | _ => NONE)
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    51
        |> the_default state
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    52
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    53
(* Too large values are understood as milliseconds, ensuring compatibility with
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    54
   old setting files. No users can possibly in their right mind want the user
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    55
   interface to block for more than 100 seconds. *)
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    56
fun smart_seconds r =
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    57
  seconds (if r >= 100.0 then
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    58
             (legacy_feature (quote auto_tools_time_limitN ^
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    59
                " expressed in milliseconds -- use seconds instead"); 0.001 * r)
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    60
           else
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    61
             r)
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    62
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    63
val _ = Context.>> (Specification.add_theorem_hook (fn interact => fn state =>
40931
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    64
  if interact andalso not (!Toplevel.quiet) andalso !time_limit > 0.0 then
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    65
    TimeLimit.timeLimit (smart_seconds (!time_limit))
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    66
        (run_tools (Data.get (Proof.theory_of state))) state
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    67
    handle TimeLimit.TimeOut => state
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    68
  else
061b8257ab9f run synchronous Auto Tools in parallel
blanchet
parents: 39330
diff changeset
    69
    state))
33561
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    70
ab01b72715ef introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
diff changeset
    71
end;