| author | wenzelm | 
| Tue, 07 Nov 2006 19:40:56 +0100 | |
| changeset 21234 | fb84ab52f23b | 
| parent 20414 | 029c4f9dc6f4 | 
| child 21254 | d53f76357f41 | 
| permissions | -rw-r--r-- | 
| 17602 | 1 | (* Title: HOL/Main.thy | 
| 2 | ID: $Id$ | |
| 3 | *) | |
| 9619 | 4 | |
| 12024 | 5 | header {* Main HOL *}
 | 
| 6 | ||
| 15131 | 7 | theory Main | 
| 17905 
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
 mengj parents: 
17721diff
changeset | 8 | imports SAT Reconstruction ResAtpMethods | 
| 15131 | 9 | begin | 
| 9650 
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
 wenzelm parents: 
9619diff
changeset | 10 | |
| 12024 | 11 | text {*
 | 
| 12 |   Theory @{text Main} includes everything.  Note that theory @{text
 | |
| 13 | PreList} already includes most HOL theories. | |
| 19608 | 14 | *} | 
| 17395 
a05e20f6a31a
hide the rather generic names used in theory Commutative_Ring;
 wenzelm parents: 
17386diff
changeset | 15 | |
| 20414 
029c4f9dc6f4
replaced skolem declarations by automatic skolemization of everything
 paulson parents: 
20362diff
changeset | 16 | text {* \medskip Late clause setup: installs \emph{all} known theorems
 | 
| 
029c4f9dc6f4
replaced skolem declarations by automatic skolemization of everything
 paulson parents: 
20362diff
changeset | 17 |   into the clause cache; cf.\ theory @{text Reconstruction}. *}
 | 
| 14350 | 18 | |
| 18510 
0a6c24f549c3
the "skolem" attribute and better initialization of the clause database
 paulson parents: 
18315diff
changeset | 19 | setup ResAxioms.setup | 
| 14350 | 20 | |
| 9650 
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
 wenzelm parents: 
9619diff
changeset | 21 | end |