| author | wenzelm |
| Tue, 31 May 2005 11:53:14 +0200 | |
| changeset 16123 | 1381e90c2694 |
| parent 16090 | fbb5ae140535 |
| child 16417 | 9bc16273c2d4 |
| permissions | -rw-r--r-- |
| 15151 | 1 |
(* Title: HOL/Reconstruction.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson |
|
4 |
Copyright 2004 University of Cambridge |
|
5 |
*) |
|
6 |
||
7 |
header{*Attributes for Reconstructing External Resolution Proofs*}
|
|
8 |
||
9 |
theory Reconstruction |
|
| 15872 | 10 |
imports Hilbert_Choice Map Infinite_Set Extraction |
|
15359
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15151
diff
changeset
|
11 |
files "Tools/res_lib.ML" |
|
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15151
diff
changeset
|
12 |
"Tools/res_clause.ML" |
|
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15151
diff
changeset
|
13 |
"Tools/res_skolem_function.ML" |
|
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15151
diff
changeset
|
14 |
"Tools/res_axioms.ML" |
|
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15151
diff
changeset
|
15 |
"Tools/res_types_sorts.ML" |
| 15653 | 16 |
|
| 16090 | 17 |
"Tools/ATP/recon_prelim.ML" |
| 15678 | 18 |
"Tools/ATP/recon_order_clauses.ML" |
19 |
"Tools/ATP/recon_translate_proof.ML" |
|
20 |
"Tools/ATP/recon_parse.ML" |
|
21 |
"Tools/ATP/recon_transfer_proof.ML" |
|
22 |
"Tools/ATP/VampireCommunication.ML" |
|
23 |
"Tools/ATP/SpassCommunication.ML" |
|
24 |
"Tools/ATP/watcher.sig" |
|
25 |
"Tools/ATP/watcher.ML" |
|
| 15700 | 26 |
"Tools/ATP/res_clasimpset.ML" |
| 15678 | 27 |
"Tools/res_atp.ML" |
28 |
||
| 16090 | 29 |
"Tools/reconstruction.ML" |
| 15151 | 30 |
|
31 |
begin |
|
32 |
||
| 15382 | 33 |
text{*Every theory of HOL must be a descendant or ancestor of this one!*}
|
34 |
||
| 16009 | 35 |
setup ResAxioms.setup |
| 15151 | 36 |
setup Reconstruction.setup |
37 |
||
| 15653 | 38 |
end |