src/HOL/Reconstruction.thy
author paulson
Tue, 12 Apr 2005 11:08:25 +0200
changeset 15700 970e0293dfb3
parent 15684 5ec4d21889d6
child 15774 9df37a0e935d
permissions -rw-r--r--
tweaks mainly to achieve sml/nj compatibility
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Reconstruction.thy
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     2
    ID: $Id$
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     4
    Copyright   2004  University of Cambridge
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     5
*)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     6
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     7
header{*Attributes for Reconstructing External Resolution Proofs*}
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     8
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     9
theory Reconstruction
15382
e56ce5cefe9c all theories must be related to Reconstruction
paulson
parents: 15359
diff changeset
    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"
15653
3549ff7158f3 Updated to add watcher code.
quigley
parents: 15645
diff changeset
    16
15678
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    17
          "Tools/ATP/recon_prelim.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    18
	  "Tools/ATP/recon_gandalf_base.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    19
 	  "Tools/ATP/recon_order_clauses.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    20
 	  "Tools/ATP/recon_translate_proof.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    21
 	  "Tools/ATP/recon_parse.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    22
 	  "Tools/ATP/recon_transfer_proof.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    23
	  "Tools/ATP/VampireCommunication.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    24
	  "Tools/ATP/SpassCommunication.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    25
	  "Tools/ATP/modUnix.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    26
	  "Tools/ATP/watcher.sig"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    27
	  "Tools/ATP/watcher.ML"
15700
970e0293dfb3 tweaks mainly to achieve sml/nj compatibility
paulson
parents: 15684
diff changeset
    28
	  "Tools/ATP/res_clasimpset.ML"
15678
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    29
	  "Tools/res_atp.ML"
28cc2314c7ff Reconstruction.thy and IsaMakefile updated
quigley
parents: 15653
diff changeset
    30
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
    31
          "Tools/reconstruction.ML"
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    32
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    33
begin
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    34
15382
e56ce5cefe9c all theories must be related to Reconstruction
paulson
parents: 15359
diff changeset
    35
text{*Every theory of HOL must be a descendant or ancestor of this one!*}
e56ce5cefe9c all theories must be related to Reconstruction
paulson
parents: 15359
diff changeset
    36
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    37
setup Reconstruction.setup
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    38
15653
3549ff7158f3 Updated to add watcher code.
quigley
parents: 15645
diff changeset
    39
end