src/Tools/solve_direct.ML
author wenzelm
Wed, 12 Apr 2017 23:44:33 +0200
changeset 65475 4519c8cc4bec
parent 64985 69592ac04836
child 67149 e61557884799
permissions -rw-r--r--
proper bootstrap base for building Pure;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
40116
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
     1
(*  Title:      Tools/solve_direct.ML
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
     2
    Author:     Timothy Bourke and Gerwin Klein, NICTA
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
     3
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
     4
Check whether a newly stated theorem can be solved directly by an
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
     5
existing theorem.  Duplicate lemmas can be detected in this way.
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
     6
39329
0a85f960ac50 make Auto Solve part of the "Auto Tools"
blanchet
parents: 39138
diff changeset
     7
The implementation relies critically on the Find_Theorems solves
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
     8
feature.
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
     9
*)
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    10
40116
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    11
signature SOLVE_DIRECT =
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    12
sig
43020
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    13
  val solve_directN: string
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    14
  val someN: string
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    15
  val noneN: string
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    16
  val unknownN: string
64012
789f5419926a modernized option
haftmann
parents: 63728
diff changeset
    17
  val max_solutions: int Config.T
64013
048b7dbfdfa3 option to report results of solve_direct as explicit warnings
haftmann
parents: 64012
diff changeset
    18
  val strict_warnings: bool Config.T
58892
20aa19ecf2cc eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents: 58843
diff changeset
    19
  val solve_direct: Proof.state -> bool * (string * string list)
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    20
end;
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    21
40130
db6e84082695 export main ML entry by default;
wenzelm
parents: 40129
diff changeset
    22
structure Solve_Direct: SOLVE_DIRECT =
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    23
struct
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    24
43025
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
    25
datatype mode = Auto_Try | Try | Normal
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
    26
40129
0843888de43e observe Isabelle/ML coding standards;
wenzelm
parents: 40116
diff changeset
    27
val solve_directN = "solve_direct";
0843888de43e observe Isabelle/ML coding standards;
wenzelm
parents: 40116
diff changeset
    28
43020
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    29
val someN = "some";
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    30
val noneN = "none";
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    31
val unknownN = "none";
40116
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    32
52701
51dfdcd88e84 guard unify tracing via visible status of global theory;
wenzelm
parents: 52645
diff changeset
    33
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    34
(* preferences *)
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    35
64012
789f5419926a modernized option
haftmann
parents: 63728
diff changeset
    36
val max_solutions = Attrib.setup_config_int @{binding solve_direct_max_solutions} (K 5);
64013
048b7dbfdfa3 option to report results of solve_direct as explicit warnings
haftmann
parents: 64012
diff changeset
    37
val strict_warnings = Attrib.setup_config_bool @{binding solve_direct_strict_warnings} (K false);
33889
4328de748fb2 some rearangement of load order to keep preferences adjacent -- slightly fragile;
wenzelm
parents: 33301
diff changeset
    38
40129
0843888de43e observe Isabelle/ML coding standards;
wenzelm
parents: 40116
diff changeset
    39
0843888de43e observe Isabelle/ML coding standards;
wenzelm
parents: 40116
diff changeset
    40
(* solve_direct command *)
0843888de43e observe Isabelle/ML coding standards;
wenzelm
parents: 40116
diff changeset
    41
43025
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
    42
fun do_solve_direct mode state =
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    43
  let
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    44
    val ctxt = Proof.context_of state;
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    45
64982
wenzelm
parents: 64013
diff changeset
    46
    fun find goal =
64983
481b2855ee9a suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
wenzelm
parents: 64982
diff changeset
    47
      #2 (Find_Theorems.find_theorems ctxt (SOME goal)
481b2855ee9a suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
wenzelm
parents: 64982
diff changeset
    48
        (SOME (Config.get ctxt max_solutions)) false [(true, Find_Theorems.Solves)]);
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    49
64985
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    50
    fun prt_result (subgoal, results) =
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    51
      Pretty.big_list
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    52
        ((if mode = Auto_Try then "Auto " else "") ^ solve_directN ^ ": " ^
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    53
          (case subgoal of NONE => "the current goal" | SOME i => "subgoal #" ^ string_of_int i) ^
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    54
          " can be solved directly with")
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    55
        (map (Find_Theorems.pretty_thm ctxt) results);
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    56
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    57
    fun seek_against_goal () =
64984
wenzelm
parents: 64983
diff changeset
    58
      (case try (#goal o Proof.simple_goal) state of
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    59
        NONE => NONE
64984
wenzelm
parents: 64983
diff changeset
    60
      | SOME goal =>
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    61
          let
64982
wenzelm
parents: 64013
diff changeset
    62
            val subgoals = Drule.strip_imp_prems (Thm.cprop_of goal);
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    63
            val rs =
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    64
              if length subgoals = 1
64982
wenzelm
parents: 64013
diff changeset
    65
              then [(NONE, find goal)]
64985
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    66
              else map_index (fn (i, sg) => (SOME (i + 1), find (Goal.init sg))) subgoals;
64982
wenzelm
parents: 64013
diff changeset
    67
            val results = filter_out (null o #2) rs;
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    68
          in if null results then NONE else SOME results end);
40116
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    69
  in
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    70
    (case try seek_against_goal () of
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    71
      SOME (SOME results) =>
43020
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    72
        (someN,
64013
048b7dbfdfa3 option to report results of solve_direct as explicit warnings
haftmann
parents: 64012
diff changeset
    73
          let
64985
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    74
            val chunks = separate (Pretty.str "") (map prt_result results);
69592ac04836 clarified message: print subgoal number instead of potentially large term;
wenzelm
parents: 64984
diff changeset
    75
            val msg = Pretty.string_of (Pretty.chunks chunks);
64013
048b7dbfdfa3 option to report results of solve_direct as explicit warnings
haftmann
parents: 64012
diff changeset
    76
          in
64982
wenzelm
parents: 64013
diff changeset
    77
            if Config.get ctxt strict_warnings then (warning msg; [])
wenzelm
parents: 64013
diff changeset
    78
            else if mode = Auto_Try then [msg]
wenzelm
parents: 64013
diff changeset
    79
            else (writeln msg; [])
64013
048b7dbfdfa3 option to report results of solve_direct as explicit warnings
haftmann
parents: 64012
diff changeset
    80
          end)
43025
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
    81
    | SOME NONE =>
64982
wenzelm
parents: 64013
diff changeset
    82
        (if mode = Normal then writeln "No proof found" else ();
58892
20aa19ecf2cc eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents: 58843
diff changeset
    83
         (noneN, []))
43025
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
    84
    | NONE =>
64982
wenzelm
parents: 64013
diff changeset
    85
        (if mode = Normal then writeln "An error occurred" else ();
58892
20aa19ecf2cc eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents: 58843
diff changeset
    86
         (unknownN, [])))
43020
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    87
  end
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    88
  |> `(fn (outcome_code, _) => outcome_code = someN);
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
    89
43025
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
    90
val solve_direct = do_solve_direct Normal
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
    91
40116
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    92
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59184
diff changeset
    93
  Outer_Syntax.command @{command_keyword solve_direct}
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 45666
diff changeset
    94
    "try to solve conjectures directly with existing theorems"
60190
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60094
diff changeset
    95
    (Scan.succeed (Toplevel.keep_proof (ignore o solve_direct o Toplevel.proof_of)));
40129
0843888de43e observe Isabelle/ML coding standards;
wenzelm
parents: 40116
diff changeset
    96
40116
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    97
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    98
(* hook *)
9ed3711366c8 introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents: 39616
diff changeset
    99
43025
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
   100
fun try_solve_direct auto = do_solve_direct (if auto then Auto_Try else Try)
5a0dec7bc099 handle non-auto try case gracefully in Solve Direct
blanchet
parents: 43024
diff changeset
   101
64982
wenzelm
parents: 64013
diff changeset
   102
val _ =
wenzelm
parents: 64013
diff changeset
   103
  Try.tool_setup (solve_directN, (10, @{system_option auto_solve_direct}, try_solve_direct));
30980
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
   104
fe0855471964 misc cleanup of auto_solve and quickcheck:
wenzelm
parents:
diff changeset
   105
end;