| author | haftmann | 
| Mon, 06 Feb 2006 11:01:28 +0100 | |
| changeset 18928 | 042608ffa2ec | 
| parent 18794 | 46d66332bf81 | 
| child 19767 | 6e77bd331bf4 | 
| permissions | -rw-r--r-- | 
| 15151 | 1 | (* Title: HOL/Reconstruction.thy | 
| 17603 | 2 | ID: $Id$ | 
| 15151 | 3 | Author: Lawrence C Paulson | 
| 4 | Copyright 2004 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 17462 
47f7bddc3239
moved setup ResAxioms.clause_setup to Main.thy (it refers to all previous theories);
 wenzelm parents: 
17421diff
changeset | 7 | header{* Reconstructing external resolution proofs *}
 | 
| 15151 | 8 | |
| 9 | theory Reconstruction | |
| 17508 | 10 | imports Hilbert_Choice Map Infinite_Set Extraction | 
| 18449 | 11 | uses "Tools/polyhash.ML" | 
| 12 | "Tools/res_clause.ML" | |
| 13 | "Tools/res_hol_clause.ML" | |
| 17819 | 14 | "Tools/res_axioms.ML" | 
| 15 | "Tools/ATP/recon_order_clauses.ML" | |
| 16 | "Tools/ATP/recon_translate_proof.ML" | |
| 17 | "Tools/ATP/recon_parse.ML" | |
| 18 | "Tools/ATP/recon_transfer_proof.ML" | |
| 19 | "Tools/ATP/AtpCommunication.ML" | |
| 20 | "Tools/ATP/watcher.ML" | |
| 18794 | 21 | "Tools/ATP/reduce_axiomsN.ML" | 
| 17819 | 22 | "Tools/ATP/res_clasimpset.ML" | 
| 23 | "Tools/res_atp.ML" | |
| 24 | "Tools/reconstruction.ML" | |
| 15151 | 25 | |
| 26 | begin | |
| 27 | ||
| 16563 | 28 | setup ResAxioms.meson_method_setup | 
| 15151 | 29 | setup Reconstruction.setup | 
| 30 | ||
| 15653 | 31 | end |