author | paulson |
Wed, 09 Feb 2005 10:17:09 +0100 | |
changeset 15508 | c09defa4c956 |
parent 15382 | e56ce5cefe9c |
child 15645 | 5e20c54683d3 |
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 |
|
15382 | 10 |
imports Hilbert_Choice Map Infinite_Set |
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" |
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15151
diff
changeset
|
16 |
"Tools/res_atp.ML" |
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15151
diff
changeset
|
17 |
"Tools/reconstruction.ML" |
15151 | 18 |
|
19 |
begin |
|
20 |
||
15382 | 21 |
text{*Every theory of HOL must be a descendant or ancestor of this one!*} |
22 |
||
15151 | 23 |
setup Reconstruction.setup |
24 |
||
25 |
end |