| author | paulson | 
| Thu, 05 Oct 2006 10:46:26 +0200 | |
| changeset 20864 | bb75b876b260 | 
| parent 17721 | b943c01e1c6d | 
| child 21332 | 2605e1ccd9f2 | 
| permissions | -rw-r--r-- | 
| 14350 | 1 | (* Title: HOL/Refute.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tjark Weber | |
| 16870 | 4 | Copyright 2003-2005 | 
| 14350 | 5 | |
| 6 | Basic setup and documentation for the 'refute' (and 'refute_params') command. | |
| 7 | *) | |
| 8 | ||
| 14589 | 9 | header {* Refute *}
 | 
| 10 | ||
| 15131 | 11 | theory Refute | 
| 15140 | 12 | imports Map | 
| 16417 | 13 | uses "Tools/prop_logic.ML" | 
| 16870 | 14 | "Tools/sat_solver.ML" | 
| 15 | "Tools/refute.ML" | |
| 16 | "Tools/refute_isar.ML" | |
| 15131 | 17 | begin | 
| 14589 | 18 | |
| 19 | setup Refute.setup | |
| 20 | ||
| 21 | text {*
 | |
| 22 | \small | |
| 23 | \begin{verbatim}
 | |
| 14350 | 24 | (* ------------------------------------------------------------------------- *) | 
| 25 | (* REFUTE *) | |
| 26 | (* *) | |
| 27 | (* We use a SAT solver to search for a (finite) model that refutes a given *) | |
| 28 | (* HOL formula. *) | |
| 29 | (* ------------------------------------------------------------------------- *) | |
| 30 | ||
| 31 | (* ------------------------------------------------------------------------- *) | |
| 14457 | 32 | (* NOTE *) | 
| 14350 | 33 | (* *) | 
| 14457 | 34 | (* I strongly recommend that you install a stand-alone SAT solver if you *) | 
| 14463 | 35 | (* want to use 'refute'. For details see 'HOL/Tools/sat_solver.ML'. If you *) | 
| 15293 
7797a04cc188
removed explicit mentioning of zChaffs version number
 webertj parents: 
15140diff
changeset | 36 | (* have installed (a supported version of) zChaff, simply set 'ZCHAFF_HOME' *) | 
| 
7797a04cc188
removed explicit mentioning of zChaffs version number
 webertj parents: 
15140diff
changeset | 37 | (* in 'etc/settings'. *) | 
| 14350 | 38 | (* ------------------------------------------------------------------------- *) | 
| 39 | ||
| 40 | (* ------------------------------------------------------------------------- *) | |
| 41 | (* USAGE *) | |
| 42 | (* *) | |
| 43 | (* See the file 'HOL/ex/Refute_Examples.thy' for examples. The supported *) | |
| 44 | (* parameters are explained below. *) | |
| 45 | (* ------------------------------------------------------------------------- *) | |
| 46 | ||
| 47 | (* ------------------------------------------------------------------------- *) | |
| 48 | (* CURRENT LIMITATIONS *) | |
| 49 | (* *) | |
| 50 | (* 'refute' currently accepts formulas of higher-order predicate logic (with *) | |
| 51 | (* equality), including free/bound/schematic variables, lambda abstractions, *) | |
| 16870 | 52 | (* sets and set membership, "arbitrary", "The", "Eps", records and *) | 
| 53 | (* inductively defined sets. Defining equations for constants are added *) | |
| 54 | (* automatically, as are sort axioms. Other, user-asserted axioms however *) | |
| 55 | (* are ignored. Inductive datatypes and recursive functions are supported, *) | |
| 56 | (* but may lead to spurious countermodels. *) | |
| 14463 | 57 | (* *) | 
| 14808 | 58 | (* The (space) complexity of the algorithm is non-elementary. *) | 
| 14350 | 59 | (* *) | 
| 16870 | 60 | (* Schematic type variables are not supported. *) | 
| 14350 | 61 | (* ------------------------------------------------------------------------- *) | 
| 62 | ||
| 63 | (* ------------------------------------------------------------------------- *) | |
| 64 | (* PARAMETERS *) | |
| 65 | (* *) | |
| 66 | (* The following global parameters are currently supported (and required): *) | |
| 67 | (* *) | |
| 68 | (* Name Type Description *) | |
| 69 | (* *) | |
| 70 | (* "minsize" int Only search for models with size at least *) | |
| 71 | (* 'minsize'. *) | |
| 72 | (* "maxsize" int If >0, only search for models with size at most *) | |
| 73 | (* 'maxsize'. *) | |
| 74 | (* "maxvars" int If >0, use at most 'maxvars' boolean variables *) | |
| 75 | (* when transforming the term into a propositional *) | |
| 76 | (* formula. *) | |
| 14808 | 77 | (* "maxtime" int If >0, terminate after at most 'maxtime' seconds. *) | 
| 78 | (* This value is ignored under some ML compilers. *) | |
| 14457 | 79 | (* "satsolver" string Name of the SAT solver to be used. *) | 
| 80 | (* *) | |
| 17721 | 81 | (* See 'HOL/SAT.thy' for default values. *) | 
| 14808 | 82 | (* *) | 
| 83 | (* The size of particular types can be specified in the form type=size *) | |
| 84 | (* (where 'type' is a string, and 'size' is an int). Examples: *) | |
| 85 | (* "'a"=1 *) | |
| 86 | (* "List.list"=2 *) | |
| 14350 | 87 | (* ------------------------------------------------------------------------- *) | 
| 88 | ||
| 89 | (* ------------------------------------------------------------------------- *) | |
| 90 | (* FILES *) | |
| 91 | (* *) | |
| 14457 | 92 | (* HOL/Tools/prop_logic.ML Propositional logic *) | 
| 93 | (* HOL/Tools/sat_solver.ML SAT solvers *) | |
| 94 | (* HOL/Tools/refute.ML Translation HOL -> propositional logic and *) | |
| 14808 | 95 | (* Boolean assignment -> HOL model *) | 
| 14350 | 96 | (* HOL/Tools/refute_isar.ML Adds 'refute'/'refute_params' to Isabelle's *) | 
| 14457 | 97 | (* syntax *) | 
| 98 | (* HOL/Refute.thy This file: loads the ML files, basic setup, *) | |
| 99 | (* documentation *) | |
| 17721 | 100 | (* HOL/SAT.thy Sets default parameters *) | 
| 14457 | 101 | (* HOL/ex/RefuteExamples.thy Examples *) | 
| 14350 | 102 | (* ------------------------------------------------------------------------- *) | 
| 14589 | 103 | \end{verbatim}
 | 
| 104 | *} | |
| 14350 | 105 | |
| 106 | end |