wenzelm@50530
|
1 |
(* Title: HOL/Library/Refute.thy
|
webertj@14350
|
2 |
Author: Tjark Weber
|
webertj@22058
|
3 |
Copyright 2003-2007
|
webertj@14350
|
4 |
|
webertj@14350
|
5 |
Basic setup and documentation for the 'refute' (and 'refute_params') command.
|
webertj@14350
|
6 |
*)
|
webertj@14350
|
7 |
|
wenzelm@14589
|
8 |
header {* Refute *}
|
wenzelm@14589
|
9 |
|
nipkow@15131
|
10 |
theory Refute
|
blanchet@40178
|
11 |
imports Hilbert_Choice List Sledgehammer
|
wenzelm@46950
|
12 |
keywords "refute" :: diag and "refute_params" :: thy_decl
|
nipkow@15131
|
13 |
begin
|
wenzelm@14589
|
14 |
|
blanchet@49985
|
15 |
ML_file "refute.ML"
|
wenzelm@14589
|
16 |
setup Refute.setup
|
wenzelm@14589
|
17 |
|
wenzelm@46960
|
18 |
refute_params
|
wenzelm@46960
|
19 |
[itself = 1,
|
wenzelm@46960
|
20 |
minsize = 1,
|
wenzelm@46960
|
21 |
maxsize = 8,
|
wenzelm@46960
|
22 |
maxvars = 10000,
|
wenzelm@46960
|
23 |
maxtime = 60,
|
wenzelm@46960
|
24 |
satsolver = auto,
|
wenzelm@46960
|
25 |
no_assms = false]
|
wenzelm@46960
|
26 |
|
wenzelm@14589
|
27 |
text {*
|
wenzelm@14589
|
28 |
\small
|
wenzelm@14589
|
29 |
\begin{verbatim}
|
webertj@14350
|
30 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
31 |
(* REFUTE *)
|
webertj@14350
|
32 |
(* *)
|
webertj@14350
|
33 |
(* We use a SAT solver to search for a (finite) model that refutes a given *)
|
webertj@14350
|
34 |
(* HOL formula. *)
|
webertj@14350
|
35 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
36 |
|
webertj@14350
|
37 |
(* ------------------------------------------------------------------------- *)
|
webertj@14457
|
38 |
(* NOTE *)
|
webertj@14350
|
39 |
(* *)
|
webertj@14457
|
40 |
(* I strongly recommend that you install a stand-alone SAT solver if you *)
|
webertj@14463
|
41 |
(* want to use 'refute'. For details see 'HOL/Tools/sat_solver.ML'. If you *)
|
webertj@15293
|
42 |
(* have installed (a supported version of) zChaff, simply set 'ZCHAFF_HOME' *)
|
webertj@15293
|
43 |
(* in 'etc/settings'. *)
|
webertj@14350
|
44 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
45 |
|
webertj@14350
|
46 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
47 |
(* USAGE *)
|
webertj@14350
|
48 |
(* *)
|
webertj@14350
|
49 |
(* See the file 'HOL/ex/Refute_Examples.thy' for examples. The supported *)
|
webertj@14350
|
50 |
(* parameters are explained below. *)
|
webertj@14350
|
51 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
52 |
|
webertj@14350
|
53 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
54 |
(* CURRENT LIMITATIONS *)
|
webertj@14350
|
55 |
(* *)
|
webertj@14350
|
56 |
(* 'refute' currently accepts formulas of higher-order predicate logic (with *)
|
webertj@14350
|
57 |
(* equality), including free/bound/schematic variables, lambda abstractions, *)
|
webertj@16870
|
58 |
(* sets and set membership, "arbitrary", "The", "Eps", records and *)
|
webertj@22058
|
59 |
(* inductively defined sets. Constants are unfolded automatically, and sort *)
|
webertj@22058
|
60 |
(* axioms are added as well. Other, user-asserted axioms however are *)
|
webertj@22058
|
61 |
(* ignored. Inductive datatypes and recursive functions are supported, but *)
|
webertj@22058
|
62 |
(* may lead to spurious countermodels. *)
|
webertj@14463
|
63 |
(* *)
|
webertj@14808
|
64 |
(* The (space) complexity of the algorithm is non-elementary. *)
|
webertj@14350
|
65 |
(* *)
|
webertj@16870
|
66 |
(* Schematic type variables are not supported. *)
|
webertj@14350
|
67 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
68 |
|
webertj@14350
|
69 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
70 |
(* PARAMETERS *)
|
webertj@14350
|
71 |
(* *)
|
blanchet@34120
|
72 |
(* The following global parameters are currently supported (and required, *)
|
blanchet@34120
|
73 |
(* except for "expect"): *)
|
webertj@14350
|
74 |
(* *)
|
webertj@14350
|
75 |
(* Name Type Description *)
|
webertj@14350
|
76 |
(* *)
|
webertj@14350
|
77 |
(* "minsize" int Only search for models with size at least *)
|
webertj@14350
|
78 |
(* 'minsize'. *)
|
webertj@14350
|
79 |
(* "maxsize" int If >0, only search for models with size at most *)
|
webertj@14350
|
80 |
(* 'maxsize'. *)
|
webertj@14350
|
81 |
(* "maxvars" int If >0, use at most 'maxvars' boolean variables *)
|
webertj@14350
|
82 |
(* when transforming the term into a propositional *)
|
webertj@14350
|
83 |
(* formula. *)
|
webertj@14808
|
84 |
(* "maxtime" int If >0, terminate after at most 'maxtime' seconds. *)
|
webertj@14808
|
85 |
(* This value is ignored under some ML compilers. *)
|
webertj@14457
|
86 |
(* "satsolver" string Name of the SAT solver to be used. *)
|
blanchet@34120
|
87 |
(* "no_assms" bool If "true", assumptions in structured proofs are *)
|
blanchet@34120
|
88 |
(* not considered. *)
|
blanchet@34120
|
89 |
(* "expect" string Expected result ("genuine", "potential", "none", or *)
|
blanchet@34120
|
90 |
(* "unknown"). *)
|
webertj@14457
|
91 |
(* *)
|
webertj@14808
|
92 |
(* The size of particular types can be specified in the form type=size *)
|
webertj@14808
|
93 |
(* (where 'type' is a string, and 'size' is an int). Examples: *)
|
webertj@14808
|
94 |
(* "'a"=1 *)
|
webertj@14808
|
95 |
(* "List.list"=2 *)
|
webertj@14350
|
96 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
97 |
|
webertj@14350
|
98 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
99 |
(* FILES *)
|
webertj@14350
|
100 |
(* *)
|
wenzelm@39048
|
101 |
(* HOL/Tools/prop_logic.ML Propositional logic *)
|
wenzelm@39048
|
102 |
(* HOL/Tools/sat_solver.ML SAT solvers *)
|
wenzelm@39048
|
103 |
(* HOL/Tools/refute.ML Translation HOL -> propositional logic and *)
|
wenzelm@39048
|
104 |
(* Boolean assignment -> HOL model *)
|
wenzelm@39048
|
105 |
(* HOL/Refute.thy This file: loads the ML files, basic setup, *)
|
wenzelm@39048
|
106 |
(* documentation *)
|
wenzelm@39048
|
107 |
(* HOL/SAT.thy Sets default parameters *)
|
wenzelm@39048
|
108 |
(* HOL/ex/Refute_Examples.thy Examples *)
|
webertj@14350
|
109 |
(* ------------------------------------------------------------------------- *)
|
wenzelm@14589
|
110 |
\end{verbatim}
|
wenzelm@14589
|
111 |
*}
|
webertj@14350
|
112 |
|
webertj@14350
|
113 |
end
|