1 (* Title: HOL/Refute.thy
6 Basic setup and documentation for the 'refute' (and 'refute_params') command.
12 imports Hilbert_Choice List Record
17 "Tools/refute_isar.ML"
25 (* ------------------------------------------------------------------------- *)
28 (* We use a SAT solver to search for a (finite) model that refutes a given *)
30 (* ------------------------------------------------------------------------- *)
32 (* ------------------------------------------------------------------------- *)
35 (* I strongly recommend that you install a stand-alone SAT solver if you *)
36 (* want to use 'refute'. For details see 'HOL/Tools/sat_solver.ML'. If you *)
37 (* have installed (a supported version of) zChaff, simply set 'ZCHAFF_HOME' *)
38 (* in 'etc/settings'. *)
39 (* ------------------------------------------------------------------------- *)
41 (* ------------------------------------------------------------------------- *)
44 (* See the file 'HOL/ex/Refute_Examples.thy' for examples. The supported *)
45 (* parameters are explained below. *)
46 (* ------------------------------------------------------------------------- *)
48 (* ------------------------------------------------------------------------- *)
49 (* CURRENT LIMITATIONS *)
51 (* 'refute' currently accepts formulas of higher-order predicate logic (with *)
52 (* equality), including free/bound/schematic variables, lambda abstractions, *)
53 (* sets and set membership, "arbitrary", "The", "Eps", records and *)
54 (* inductively defined sets. Constants are unfolded automatically, and sort *)
55 (* axioms are added as well. Other, user-asserted axioms however are *)
56 (* ignored. Inductive datatypes and recursive functions are supported, but *)
57 (* may lead to spurious countermodels. *)
59 (* The (space) complexity of the algorithm is non-elementary. *)
61 (* Schematic type variables are not supported. *)
62 (* ------------------------------------------------------------------------- *)
64 (* ------------------------------------------------------------------------- *)
67 (* The following global parameters are currently supported (and required): *)
69 (* Name Type Description *)
71 (* "minsize" int Only search for models with size at least *)
73 (* "maxsize" int If >0, only search for models with size at most *)
75 (* "maxvars" int If >0, use at most 'maxvars' boolean variables *)
76 (* when transforming the term into a propositional *)
78 (* "maxtime" int If >0, terminate after at most 'maxtime' seconds. *)
79 (* This value is ignored under some ML compilers. *)
80 (* "satsolver" string Name of the SAT solver to be used. *)
82 (* See 'HOL/SAT.thy' for default values. *)
84 (* The size of particular types can be specified in the form type=size *)
85 (* (where 'type' is a string, and 'size' is an int). Examples: *)
88 (* ------------------------------------------------------------------------- *)
90 (* ------------------------------------------------------------------------- *)
93 (* HOL/Tools/prop_logic.ML Propositional logic *)
94 (* HOL/Tools/sat_solver.ML SAT solvers *)
95 (* HOL/Tools/refute.ML Translation HOL -> propositional logic and *)
96 (* Boolean assignment -> HOL model *)
97 (* HOL/Tools/refute_isar.ML Adds 'refute'/'refute_params' to Isabelle's *)
99 (* HOL/Refute.thy This file: loads the ML files, basic setup, *)
101 (* HOL/SAT.thy Sets default parameters *)
102 (* HOL/ex/RefuteExamples.thy Examples *)
103 (* ------------------------------------------------------------------------- *)