src/HOL/Tools/res_atp.ML
author paulson
Wed, 09 Aug 2006 18:39:08 +0200
changeset 20372 e42674e0486e
parent 20289 ba7a7c56bed5
child 20419 df257a9cf0e9
permissions -rw-r--r--
blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
     1
(*  Author: Jia Meng, Cambridge University Computer Laboratory, NICTA
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
     2
    ID: $Id$
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
     3
    Copyright 2004 University of Cambridge
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     4
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     5
ATPs with TPTP format input.
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     6
*)
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
     7
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
     8
signature RES_ATP =
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
     9
sig
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
    10
  val prover: string ref
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
    11
  val custom_spass: string list ref
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    12
  val destdir: string ref
17849
d7619ccf22e6 signature changes
paulson
parents: 17845
diff changeset
    13
  val helper_path: string -> string -> string
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    14
  val problem_name: string ref
17690
8ba7c3cd24a8 time limit option; fixed bug concerning first line of ATP output
paulson
parents: 17525
diff changeset
    15
  val time_limit: int ref
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    16
   
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    17
  datatype mode = Auto | Fol | Hol
19450
651d949776f8 exported linkup_logic_mode and changed the default setting
paulson
parents: 19445
diff changeset
    18
  val linkup_logic_mode : mode ref
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    19
  val write_subgoal_file: bool -> mode -> Proof.context -> thm list -> thm list -> int -> string
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    20
  val vampire_time: int ref
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    21
  val eprover_time: int ref
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    22
  val spass_time: int ref
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    23
  val run_vampire: int -> unit
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    24
  val run_eprover: int -> unit
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    25
  val run_spass: int -> unit
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    26
  val vampireLimit: unit -> int
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    27
  val eproverLimit: unit -> int
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    28
  val spassLimit: unit -> int
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20246
diff changeset
    29
  val atp_method: (Proof.context -> thm list -> int -> tactic) ->
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20246
diff changeset
    30
    Method.src -> Proof.context -> Proof.method
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    31
  val cond_rm_tmp: string -> unit
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    32
  val keep_atp_input: bool ref
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    33
  val fol_keep_types: bool ref
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    34
  val hol_full_types: unit -> unit
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    35
  val hol_partial_types: unit -> unit
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    36
  val hol_const_types_only: unit -> unit
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    37
  val hol_no_types: unit -> unit
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    38
  val hol_typ_level: unit -> ResHolClause.type_level
20246
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
    39
  val include_all: bool ref
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    40
  val run_relevance_filter: bool ref
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
    41
  val run_blacklist_filter: bool ref
20372
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
    42
  val blacklist: string list ref
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
    43
  val add_all : unit -> unit
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    44
  val add_claset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    45
  val add_simpset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    46
  val add_clasimp : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    47
  val add_atpset : unit -> unit
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
    48
  val rm_all : unit -> unit
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    49
  val rm_claset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    50
  val rm_simpset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    51
  val rm_atpset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    52
  val rm_clasimp : unit -> unit
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    53
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    54
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
    55
structure ResAtp : RES_ATP =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    56
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    57
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    58
(********************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    59
(* some settings for both background automatic ATP calling procedure*)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    60
(* and also explicit ATP invocation methods                         *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    61
(********************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    62
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    63
(*** background linkup ***)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    64
val call_atp = ref false; 
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 17091
diff changeset
    65
val hook_count = ref 0;
18675
333a73034023 more practical time limit
paulson
parents: 18404
diff changeset
    66
val time_limit = ref 30;
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17317
diff changeset
    67
val prover = ref "E";   (* use E as the default prover *)
17305
6cef3aedd661 axioms now included in tptp files, no /bin/cat and various tidying
paulson
parents: 17235
diff changeset
    68
val custom_spass =   (*specialized options for SPASS*)
17690
8ba7c3cd24a8 time limit option; fixed bug concerning first line of ATP output
paulson
parents: 17525
diff changeset
    69
      ref ["-Auto=0","-FullRed=0","-IORe","-IOFc","-RTaut","-RFSub","-RBSub"];
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    70
val destdir = ref "";   (*Empty means write files to /tmp*)
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    71
val problem_name = ref "prob";
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    72
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    73
(*Return the path to a "helper" like SPASS or tptp2X, first checking that
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    74
  it exists.  FIXME: modify to use Path primitives and move to some central place.*)  
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    75
fun helper_path evar base =
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    76
  case getenv evar of
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    77
      "" => error  ("Isabelle environment variable " ^ evar ^ " not defined")
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    78
    | home => 
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    79
        let val path = home ^ "/" ^ base
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    80
        in  if File.exists (File.unpack_platform_path path) then path 
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    81
	    else error ("Could not find the file " ^ path)
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    82
	end;  
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    83
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
    84
fun probfile_nosuffix _ = 
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    85
  if !destdir = "" then File.platform_path (File.tmp_path (Path.basic (!problem_name)))
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    86
  else if File.exists (File.unpack_platform_path (!destdir))
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    87
  then !destdir ^ "/" ^ !problem_name
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    88
  else error ("No such directory: " ^ !destdir);
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    89
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
    90
fun prob_pathname n = probfile_nosuffix n ^ "_" ^ Int.toString n;
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
    91
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    92
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    93
(*** ATP methods ***)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    94
val vampire_time = ref 60;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    95
val eprover_time = ref 60;
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    96
val spass_time = ref 60;
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    97
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    98
fun run_vampire time =  
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    99
    if (time >0) then vampire_time:= time
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   100
    else vampire_time:=60;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   101
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   102
fun run_eprover time = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   103
    if (time > 0) then eprover_time:= time
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   104
    else eprover_time:=60;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   105
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   106
fun run_spass time = 
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   107
    if (time > 0) then spass_time:=time
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   108
    else spass_time:=60;
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   109
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   110
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   111
fun vampireLimit () = !vampire_time;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   112
fun eproverLimit () = !eprover_time;
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   113
fun spassLimit () = !spass_time;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   114
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   115
val keep_atp_input = ref false;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   116
val fol_keep_types = ResClause.keep_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   117
val hol_full_types = ResHolClause.full_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   118
val hol_partial_types = ResHolClause.partial_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   119
val hol_const_types_only = ResHolClause.const_types_only;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   120
val hol_no_types = ResHolClause.no_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   121
fun hol_typ_level () = ResHolClause.find_typ_level ();
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   122
fun is_typed_hol () = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   123
    let val tp_level = hol_typ_level()
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   124
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   125
	not (tp_level = ResHolClause.T_NONE)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   126
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   127
val include_combS = ResHolClause.include_combS;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   128
val include_min_comb = ResHolClause.include_min_comb;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   129
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   130
fun atp_input_file () =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   131
    let val file = !problem_name 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   132
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   133
	if !destdir = "" then File.platform_path (File.tmp_path (Path.basic file))
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   134
	else if File.exists (File.unpack_platform_path (!destdir))
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   135
	then !destdir ^ "/" ^ file
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   136
	else error ("No such directory: " ^ !destdir)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   137
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   138
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   139
val include_all = ref false;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   140
val include_simpset = ref false;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   141
val include_claset = ref false; 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   142
val include_atpset = ref true;
20246
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   143
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   144
(*Tests show that follow_defs gives VERY poor results with "include_all"*)
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   145
fun add_all() = (include_all:=true; ReduceAxiomsN.follow_defs := false);
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   146
fun rm_all() = include_all:=false;
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   147
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   148
fun add_simpset() = include_simpset:=true;
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   149
fun rm_simpset() = include_simpset:=false;
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   150
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   151
fun add_claset() = include_claset:=true;
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   152
fun rm_claset() = include_claset:=false;
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   153
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   154
fun add_clasimp() = (include_simpset:=true;include_claset:=true);
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   155
fun rm_clasimp() = (include_simpset:=false;include_claset:=false);
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   156
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   157
fun add_atpset() = include_atpset:=true;
fdfe7399e057 "all theorems" mode forces definition-expansion off.
paulson
parents: 20224
diff changeset
   158
fun rm_atpset() = include_atpset:=false;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   159
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   160
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   161
(**** relevance filter ****)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   162
val run_relevance_filter = ref true;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   163
val run_blacklist_filter = ref true;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   164
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   165
(******************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   166
(* detect whether a given problem (clauses) is FOL/HOL/HOLC/HOLCS *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   167
(******************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   168
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   169
datatype logic = FOL | HOL | HOLC | HOLCS;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   170
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   171
fun string_of_logic FOL = "FOL"
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   172
  | string_of_logic HOL = "HOL"
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   173
  | string_of_logic HOLC = "HOLC"
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   174
  | string_of_logic HOLCS = "HOLCS";
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   175
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   176
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   177
fun is_fol_logic FOL = true
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   178
  | is_fol_logic  _ = false
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   179
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   180
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   181
(*HOLCS will not occur here*)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   182
fun upgrade_lg HOLC _ = HOLC
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   183
  | upgrade_lg HOL HOLC = HOLC
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   184
  | upgrade_lg HOL _ = HOL
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   185
  | upgrade_lg FOL lg = lg; 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   186
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   187
(* check types *)
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   188
fun has_bool_hfn (Type("bool",_)) = true
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   189
  | has_bool_hfn (Type("fun",_)) = true
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   190
  | has_bool_hfn (Type(_, Ts)) = exists has_bool_hfn Ts
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   191
  | has_bool_hfn _ = false;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   192
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   193
fun is_hol_fn tp =
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   194
    let val (targs,tr) = strip_type tp
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   195
    in
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   196
	exists (has_bool_hfn) (tr::targs)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   197
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   198
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   199
fun is_hol_pred tp =
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   200
    let val (targs,tr) = strip_type tp
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   201
    in
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   202
	exists (has_bool_hfn) targs
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   203
    end;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   204
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   205
exception FN_LG of term;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   206
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   207
fun fn_lg (t as Const(f,tp)) (lg,seen) = 
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   208
    if is_hol_fn tp then (upgrade_lg HOL lg, t ins seen) else (lg, t ins seen) 
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   209
  | fn_lg (t as Free(f,tp)) (lg,seen) = 
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   210
    if is_hol_fn tp then (upgrade_lg HOL lg, t ins seen) else (lg, t ins seen) 
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   211
  | fn_lg (t as Var(f,tp)) (lg,seen) =
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   212
    if is_hol_fn tp then (upgrade_lg HOL lg,t ins seen) else (lg,t ins seen)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   213
  | fn_lg (t as Abs(_,_,_)) (lg,seen) = (upgrade_lg HOLC lg,t ins seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   214
  | fn_lg f _ = raise FN_LG(f); 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   215
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   216
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   217
fun term_lg [] (lg,seen) = (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   218
  | term_lg (tm::tms) (FOL,seen) =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   219
    let val (f,args) = strip_comb tm
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   220
	val (lg',seen') = if f mem seen then (FOL,seen) 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   221
			  else fn_lg f (FOL,seen)
19641
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   222
	val _ =
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   223
          if is_fol_logic lg' then ()
19746
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   224
          else Output.debug ("Found a HOL term: " ^ Display.raw_string_of_term f)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   225
	 in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   226
	     term_lg (args@tms) (lg',seen')
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   227
    end
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   228
  | term_lg _ (lg,seen) = (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   229
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   230
exception PRED_LG of term;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   231
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   232
fun pred_lg (t as Const(P,tp)) (lg,seen)= 
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   233
    if is_hol_pred tp then (upgrade_lg HOL lg, t ins seen) else (lg,t ins seen) 
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   234
  | pred_lg (t as Free(P,tp)) (lg,seen) =
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   235
    if is_hol_pred tp then (upgrade_lg HOL lg, t ins seen) else (lg,t ins seen)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   236
  | pred_lg (t as Var(_,_)) (lg,seen) = (upgrade_lg HOL lg, t ins seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   237
  | pred_lg P _ = raise PRED_LG(P);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   238
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   239
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   240
fun lit_lg (Const("Not",_) $ P) (lg,seen) = lit_lg P (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   241
  | lit_lg P (lg,seen) =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   242
    let val (pred,args) = strip_comb P
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   243
	val (lg',seen') = if pred mem seen then (lg,seen) 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   244
			  else pred_lg pred (lg,seen)
19641
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   245
	val _ =
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   246
          if is_fol_logic lg' then ()
19746
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   247
          else Output.debug ("Found a HOL predicate: " ^ Display.raw_string_of_term pred)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   248
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   249
	term_lg args (lg',seen')
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   250
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   251
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   252
fun lits_lg [] (lg,seen) = (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   253
  | lits_lg (lit::lits) (FOL,seen) =
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   254
    let val (lg,seen') = lit_lg lit (FOL,seen)
19641
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   255
	val _ =
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   256
          if is_fol_logic lg then ()
19746
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   257
          else Output.debug ("Found a HOL literal: " ^ Display.raw_string_of_term lit)
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   258
    in
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   259
	lits_lg lits (lg,seen')
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   260
    end
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   261
  | lits_lg lits (lg,seen) = (lg,seen);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   262
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   263
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   264
fun dest_disj_aux (Const ("op |", _) $ t $ t') disjs = 
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   265
    dest_disj_aux t (dest_disj_aux t' disjs)
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   266
  | dest_disj_aux t disjs = t::disjs;
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   267
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   268
fun dest_disj t = dest_disj_aux t [];
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   269
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   270
fun logic_of_clause tm (lg,seen) =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   271
    let val tm' = HOLogic.dest_Trueprop tm
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   272
	val disjs = dest_disj tm'
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   273
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   274
	lits_lg disjs (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   275
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   276
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   277
fun logic_of_clauses [] (lg,seen) = (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   278
  | logic_of_clauses (cls::clss) (FOL,seen) =
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   279
    let val (lg,seen') = logic_of_clause cls (FOL,seen)
19641
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   280
	val _ =
f1de44e61ec1 replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents: 19617
diff changeset
   281
          if is_fol_logic lg then ()
19746
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   282
          else Output.debug ("Found a HOL clause: " ^ Display.raw_string_of_term cls)
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   283
    in
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   284
	logic_of_clauses clss (lg,seen')
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   285
    end
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   286
  | logic_of_clauses (cls::clss) (lg,seen) = (lg,seen);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   287
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   288
fun problem_logic_goals_aux [] (lg,seen) = lg
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   289
  | problem_logic_goals_aux (subgoal::subgoals) (lg,seen) = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   290
    problem_logic_goals_aux subgoals (logic_of_clauses subgoal (lg,seen));
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   291
    
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   292
fun problem_logic_goals subgoals = problem_logic_goals_aux subgoals (FOL,[]);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   293
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   294
(***************************************************************)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   295
(* Retrieving and filtering lemmas                             *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   296
(***************************************************************)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   297
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   298
(*** white list and black list of lemmas ***)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   299
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   300
(*The rule subsetI is frequently omitted by the relevance filter.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   301
val whitelist = ref [subsetI]; 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   302
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   303
(*In general, these produce clauses that are prolific (match too many equality or
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   304
  membership literals) and relate to seldom-used facts. Some duplicate other rules.
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   305
  FIXME: this blacklist needs to be maintained using theory data and added to using
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   306
  an attribute.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   307
val blacklist = ref
20372
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   308
  ["Datatype.unit.split_asm",         (*These  "unit" thms cause unsound proofs*)
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   309
                               (*Datatype.unit.nchotomy is caught automatically*)
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   310
   "Datatype.unit.induct",
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   311
   "Datatype.unit.inducts",
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   312
   "Datatype.unit.split",
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   313
   "Datatype.unit.splits_1",
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   314
   "Datatype.unit.splits_2",
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   315
   "Product_Type.unit_abs_eta_conv",
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   316
   "Product_Type.unit_induct",
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   317
e42674e0486e blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
paulson
parents: 20289
diff changeset
   318
   "Datatype.prod.size",
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   319
   "Divides.dvd_0_left_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   320
   "Finite_Set.card_0_eq",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   321
   "Finite_Set.card_infinite",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   322
   "Finite_Set.Max_ge",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   323
   "Finite_Set.Max_in",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   324
   "Finite_Set.Max_le_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   325
   "Finite_Set.Max_less_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   326
   "Finite_Set.max.f_below_strict_below.below_f_conv", (*duplicates in Orderings.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   327
   "Finite_Set.max.f_below_strict_below.strict_below_f_conv", (*duplicates in Orderings.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   328
   "Finite_Set.Min_ge_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   329
   "Finite_Set.Min_gr_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   330
   "Finite_Set.Min_in",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   331
   "Finite_Set.Min_le",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   332
   "Finite_Set.min_max.below_inf_sup_Inf_Sup.inf_Sup_absorb", 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   333
   "Finite_Set.min_max.below_inf_sup_Inf_Sup.sup_Inf_absorb", 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   334
   "Finite_Set.min.f_below_strict_below.below_f_conv",        (*duplicates in Orderings.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   335
   "Finite_Set.min.f_below_strict_below.strict_below_f_conv", (*duplicates in Orderings.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   336
   "IntDef.Integ.Abs_Integ_inject",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   337
   "IntDef.Integ.Abs_Integ_inverse",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   338
   "IntDiv.zdvd_0_left",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   339
   "List.append_eq_append_conv",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   340
   "List.hd_Cons_tl",   (*Says everything is [] or Cons. Probably prolific.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   341
   "List.in_listsD",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   342
   "List.in_listsI",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   343
   "List.lists.Cons",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   344
   "List.listsE",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   345
   "Nat.less_one", (*not directional? obscure*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   346
   "Nat.not_gr0",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   347
   "Nat.one_eq_mult_iff", (*duplicate by symmetry*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   348
   "NatArith.of_nat_0_eq_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   349
   "NatArith.of_nat_eq_0_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   350
   "NatArith.of_nat_le_0_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   351
   "NatSimprocs.divide_le_0_iff_number_of",  (*too many clauses*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   352
   "NatSimprocs.divide_less_0_iff_number_of",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   353
   "NatSimprocs.equation_minus_iff_1",  (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   354
   "NatSimprocs.equation_minus_iff_number_of", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   355
   "NatSimprocs.le_minus_iff_1", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   356
   "NatSimprocs.le_minus_iff_number_of",  (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   357
   "NatSimprocs.less_minus_iff_1", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   358
   "NatSimprocs.less_minus_iff_number_of", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   359
   "NatSimprocs.minus_equation_iff_number_of", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   360
   "NatSimprocs.minus_le_iff_1", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   361
   "NatSimprocs.minus_le_iff_number_of", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   362
   "NatSimprocs.minus_less_iff_1", (*not directional*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   363
   "NatSimprocs.mult_le_cancel_left_number_of", (*excessive case analysis*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   364
   "NatSimprocs.mult_le_cancel_right_number_of", (*excessive case analysis*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   365
   "NatSimprocs.mult_less_cancel_left_number_of", (*excessive case analysis*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   366
   "NatSimprocs.mult_less_cancel_right_number_of", (*excessive case analysis*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   367
   "NatSimprocs.zero_le_divide_iff_number_of", (*excessive case analysis*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   368
   "NatSimprocs.zero_less_divide_iff_number_of",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   369
   "OrderedGroup.abs_0_eq", (*duplicate by symmetry*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   370
   "OrderedGroup.diff_eq_0_iff_eq", (*prolific?*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   371
   "OrderedGroup.join_0_eq_0",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   372
   "OrderedGroup.meet_0_eq_0",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   373
   "OrderedGroup.pprt_eq_0",   (*obscure*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   374
   "OrderedGroup.pprt_eq_id",   (*obscure*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   375
   "OrderedGroup.pprt_mono",   (*obscure*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   376
   "Parity.even_nat_power",   (*obscure, somewhat prolilfic*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   377
   "Parity.power_eq_0_iff_number_of",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   378
   "Parity.power_le_zero_eq_number_of",   (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   379
   "Parity.power_less_zero_eq_number_of",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   380
   "Parity.zero_le_power_eq_number_of",   (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   381
   "Parity.zero_less_power_eq_number_of",   (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   382
   "Power.zero_less_power_abs_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   383
   "Relation.diagI",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   384
   "Relation.ImageI",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   385
   "Ring_and_Field.divide_cancel_left", (*fields are seldom used & often prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   386
   "Ring_and_Field.divide_cancel_right",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   387
   "Ring_and_Field.divide_divide_eq_left",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   388
   "Ring_and_Field.divide_divide_eq_right",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   389
   "Ring_and_Field.divide_eq_0_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   390
   "Ring_and_Field.divide_eq_1_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   391
   "Ring_and_Field.divide_eq_eq_1",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   392
   "Ring_and_Field.divide_le_0_1_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   393
   "Ring_and_Field.divide_le_eq_1_neg",  (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   394
   "Ring_and_Field.divide_le_eq_1_pos",  (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   395
   "Ring_and_Field.divide_less_0_1_iff",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   396
   "Ring_and_Field.divide_less_eq_1_neg",  (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   397
   "Ring_and_Field.divide_less_eq_1_pos",  (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   398
   "Ring_and_Field.eq_divide_eq_1", (*duplicate by symmetry*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   399
   "Ring_and_Field.field_mult_cancel_left",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   400
   "Ring_and_Field.field_mult_cancel_right",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   401
   "Ring_and_Field.inverse_le_iff_le_neg",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   402
   "Ring_and_Field.inverse_le_iff_le",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   403
   "Ring_and_Field.inverse_less_iff_less_neg",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   404
   "Ring_and_Field.inverse_less_iff_less",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   405
   "Ring_and_Field.le_divide_eq_1_neg", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   406
   "Ring_and_Field.le_divide_eq_1_pos", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   407
   "Ring_and_Field.less_divide_eq_1_neg", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   408
   "Ring_and_Field.less_divide_eq_1_pos", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   409
   "Ring_and_Field.one_eq_divide_iff",  (*duplicate by symmetry*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   410
   "Set.Diff_eq_empty_iff", (*redundant with paramodulation*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   411
   "Set.Diff_insert0",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   412
   "Set.disjoint_insert_1",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   413
   "Set.disjoint_insert_2",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   414
   "Set.empty_Union_conv", (*redundant with paramodulation*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   415
   "Set.insert_disjoint_1",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   416
   "Set.insert_disjoint_2",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   417
   "Set.Int_UNIV", (*redundant with paramodulation*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   418
   "Set.Inter_iff",              (*We already have InterI, InterE*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   419
   "Set.Inter_UNIV_conv_1",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   420
   "Set.Inter_UNIV_conv_2",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   421
   "Set.psubsetE",    (*too prolific and obscure*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   422
   "Set.psubsetI",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   423
   "Set.singleton_insert_inj_eq'",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   424
   "Set.singleton_insert_inj_eq",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   425
   "Set.singletonD",  (*these two duplicate some "insert" lemmas*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   426
   "Set.singletonI",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   427
   "Set.Un_empty", (*redundant with paramodulation*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   428
   "Set.Union_empty_conv", (*redundant with paramodulation*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   429
   "Set.Union_iff",              (*We already have UnionI, UnionE*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   430
   "SetInterval.atLeastAtMost_iff", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   431
   "SetInterval.atLeastLessThan_iff", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   432
   "SetInterval.greaterThanAtMost_iff", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   433
   "SetInterval.greaterThanLessThan_iff", (*obscure and prolific*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   434
   "SetInterval.ivl_subset"];  (*excessive case analysis*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   435
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   436
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   437
(*These might be prolific but are probably OK, and min and max are basic.
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   438
   "Orderings.max_less_iff_conj", 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   439
   "Orderings.min_less_iff_conj",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   440
   "Orderings.min_max.below_inf.below_inf_conv",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   441
   "Orderings.min_max.below_sup.above_sup_conv",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   442
Very prolific and somewhat obscure:
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   443
   "Set.InterD",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   444
   "Set.UnionI",
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   445
*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   446
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   447
(*** retrieve lemmas from clasimpset and atpset, may filter them ***)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   448
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   449
(*The "name" of a theorem is its statement, if nothing else is available.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   450
val plain_string_of_thm =
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   451
    setmp show_question_marks false 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   452
      (setmp print_mode [] 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   453
	(Pretty.setmp_margin 999 string_of_thm));
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   454
	
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   455
(*Returns the first substring enclosed in quotation marks, typically omitting 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   456
  the [.] of meta-level assumptions.*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   457
val firstquoted = hd o (String.tokens (fn c => c = #"\""))
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   458
	
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   459
fun fake_thm_name th = 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   460
    Context.theory_name (theory_of_thm th) ^ "." ^ firstquoted (plain_string_of_thm th);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   461
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   462
fun put_name_pair ("",th) = (fake_thm_name th, th)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   463
  | put_name_pair (a,th)  = (a,th);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   464
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   465
(*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   466
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   467
exception HASH_CLAUSE and HASH_STRING;
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   468
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   469
(*Catches (for deletion) theorems automatically generated from other theorems*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   470
fun insert_suffixed_names ht x = 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   471
     (Polyhash.insert ht (x^"_iff1", ()); 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   472
      Polyhash.insert ht (x^"_iff2", ()); 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   473
      Polyhash.insert ht (x^"_dest", ())); 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   474
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   475
fun make_banned_test xs = 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   476
  let val ht = Polyhash.mkTable (Polyhash.hash_string, op =)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   477
                                (6000, HASH_STRING)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   478
      fun banned s = isSome (Polyhash.peek ht s)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   479
  in  app (fn x => Polyhash.insert ht (x,())) (!blacklist);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   480
      app (insert_suffixed_names ht) (!blacklist @ xs); 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   481
      banned
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   482
  end;
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   483
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   484
(** a hash function from Term.term to int, and also a hash table **)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   485
val xor_words = List.foldl Word.xorb 0w0;
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   486
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   487
fun hashw_term ((Const(c,_)), w) = Polyhash.hashw_string (c,w)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   488
  | hashw_term ((Free(_,_)), w) = w
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   489
  | hashw_term ((Var(_,_)), w) = w
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   490
  | hashw_term ((Bound _), w) = w
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   491
  | hashw_term ((Abs(_,_,t)), w) = hashw_term (t, w)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   492
  | hashw_term ((P$Q), w) = hashw_term (Q, (hashw_term (P, w)));
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   493
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   494
fun hashw_pred (P,w) = 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   495
    let val (p,args) = strip_comb P
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   496
    in
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   497
	List.foldl hashw_term w (p::args)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   498
    end;
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   499
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   500
fun hash_literal (Const("Not",_)$P) = Word.notb(hashw_pred(P,0w0))
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   501
  | hash_literal P = hashw_pred(P,0w0);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   502
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   503
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   504
fun get_literals (Const("Trueprop",_)$P) lits = get_literals P lits
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   505
  | get_literals (Const("op |",_)$P$Q) lits = get_literals Q (get_literals P lits)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   506
  | get_literals lit lits = (lit::lits);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   507
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   508
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   509
fun hash_term term = Word.toIntX (xor_words (map hash_literal (get_literals term [])));
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   510
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   511
fun hash_thm  thm = hash_term (prop_of thm);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   512
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   513
fun equal_thm (thm1,thm2) = Term.aconv(prop_of thm1, prop_of thm2);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   514
(*Create a hash table for clauses, of the given size*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   515
fun mk_clause_table n =
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   516
      Polyhash.mkTable (hash_thm, equal_thm)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   517
                       (n, HASH_CLAUSE);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   518
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   519
(*Use a hash table to eliminate duplicates from xs*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   520
fun make_unique ht xs = 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   521
      (app (ignore o Polyhash.peekInsert ht) xs;  Polyhash.listItems ht);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   522
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   523
fun mem_thm th [] = false
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   524
  | mem_thm th ((th',_)::thms_names) = equal_thm (th,th') orelse mem_thm th thms_names;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   525
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   526
fun insert_thms [] thms_names = thms_names
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   527
  | insert_thms ((thm,name)::thms_names) thms_names' =
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   528
      if mem_thm thm thms_names' then insert_thms thms_names thms_names' 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   529
      else insert_thms thms_names ((thm,name)::thms_names');
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   530
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   531
fun display_thms [] = ()
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   532
  | display_thms ((name,thm)::nthms) = 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   533
      let val nthm = name ^ ": " ^ (string_of_thm thm)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   534
      in Output.debug nthm; display_thms nthms  end;
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   535
 
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   536
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   537
fun all_facts_of ctxt =
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   538
  FactIndex.find (ProofContext.fact_index_of ctxt) ([], [])
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   539
  |> maps #2 |> map (`Thm.name_of_thm);
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   540
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   541
(* get lemmas from claset, simpset, atpset and extra supplied rules *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   542
fun get_clasimp_atp_lemmas ctxt user_thms = 
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   543
  let val included_thms =
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   544
	if !include_all 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   545
	then (tap (fn ths => Output.debug ("Including all " ^ Int.toString (length ths) ^
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   546
	                                   " theorems")) 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   547
	          (all_facts_of ctxt @ PureThy.all_thms_of (ProofContext.theory_of ctxt)))
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   548
	else 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   549
	let val claset_thms =
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   550
		if !include_claset then ResAxioms.claset_rules_of_ctxt ctxt
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   551
		else []
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   552
	    val simpset_thms = 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   553
		if !include_simpset then ResAxioms.simpset_rules_of_ctxt ctxt
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   554
		else []
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   555
	    val atpset_thms =
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   556
		if !include_atpset then ResAxioms.atpset_rules_of_ctxt ctxt
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   557
		else []
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   558
	    val _ = if !Output.show_debug_msgs 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   559
		    then (Output.debug "ATP theorems: "; display_thms atpset_thms) 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   560
		    else ()		 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   561
	in  claset_thms @ simpset_thms @ atpset_thms  end
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   562
      val user_rules = map (put_name_pair o ResAxioms.pairname)
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   563
			   (if null user_thms then !whitelist else user_thms)
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   564
  in
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   565
      (map put_name_pair included_thms, user_rules)
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   566
  end;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   567
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   568
(* remove lemmas that are banned from the backlist *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   569
fun blacklist_filter thms = 
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   570
  if !run_blacklist_filter then 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   571
      let val banned = make_banned_test (map #1 thms)
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   572
	  fun ok (a,_) = not (banned a)
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   573
      in  filter ok thms  end
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   574
  else thms;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   575
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   576
(* filter axiom clauses, but keep supplied clauses and clauses in whitelist *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   577
fun get_relevant_clauses ctxt cls_thms white_cls goals =
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   578
 let val cls_thms_list = make_unique (mk_clause_table 2200) (List.concat (white_cls@cls_thms))
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   579
     val relevant_cls_thms_list = 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   580
	 if !run_relevance_filter 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   581
	 then ReduceAxiomsN.relevance_filter (ProofContext.theory_of ctxt) cls_thms_list goals 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   582
	 else cls_thms_list
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   583
 in
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   584
     insert_thms (List.concat white_cls) relevant_cls_thms_list 
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   585
 end;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   586
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   587
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   588
(* ATP invocation methods setup                                *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   589
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   590
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   591
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   592
(**** prover-specific format: TPTP ****)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   593
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   594
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   595
fun cnf_hyps_thms ctxt = 
20224
9c40a144ee0e moved basic assumption operations from structure ProofContext to Assumption;
wenzelm
parents: 20131
diff changeset
   596
    let val ths = Assumption.prems_of ctxt
19617
7cb4b67d4b97 avoid raw equality on type thm;
wenzelm
parents: 19490
diff changeset
   597
    in fold (fold (insert Thm.eq_thm) o ResAxioms.skolem_thm) ths [] end;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   598
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   599
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   600
(**** write to files ****)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   601
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   602
datatype mode = Auto | Fol | Hol;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   603
19450
651d949776f8 exported linkup_logic_mode and changed the default setting
paulson
parents: 19445
diff changeset
   604
val linkup_logic_mode = ref Auto;
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   605
20131
c89ee2f4efd5 Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
mengj
parents: 20022
diff changeset
   606
fun tptp_writer logic goals filename (axioms,classrels,arities) user_lemmas =
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   607
    if is_fol_logic logic 
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   608
    then ResClause.tptp_write_file goals filename (axioms, classrels, arities)
20131
c89ee2f4efd5 Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
mengj
parents: 20022
diff changeset
   609
    else ResHolClause.tptp_write_file goals filename (axioms, classrels, arities) user_lemmas;
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   610
20131
c89ee2f4efd5 Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
mengj
parents: 20022
diff changeset
   611
fun dfg_writer logic goals filename (axioms,classrels,arities) user_lemmas =
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   612
    if is_fol_logic logic 
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   613
    then ResClause.dfg_write_file goals filename (axioms, classrels, arities)
20131
c89ee2f4efd5 Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
mengj
parents: 20022
diff changeset
   614
    else ResHolClause.dfg_write_file goals filename (axioms, classrels, arities) user_lemmas;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   615
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   616
(*Called by the oracle-based methods declared in res_atp_methods.ML*)
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   617
fun write_subgoal_file dfg mode ctxt conjectures user_thms n =
19442
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   618
    let val conj_cls = make_clauses conjectures 
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   619
	val hyp_cls = cnf_hyps_thms ctxt
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   620
	val goal_cls = conj_cls@hyp_cls
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   621
	val (included_thms,user_rules) = get_clasimp_atp_lemmas ctxt user_thms
20131
c89ee2f4efd5 Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
mengj
parents: 20022
diff changeset
   622
	val user_lemmas_names = map #1 user_rules
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   623
	val rm_black_cls = blacklist_filter included_thms 
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   624
	val cla_simp_atp_clauses = ResAxioms.cnf_rules_pairs rm_black_cls
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   625
	val user_cls = ResAxioms.cnf_rules_pairs user_rules
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   626
	val axclauses = get_relevant_clauses ctxt cla_simp_atp_clauses
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   627
	                            user_cls (map prop_of goal_cls)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   628
	val thy = ProofContext.theory_of ctxt
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   629
	val prob_logic = case mode of 
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   630
                            Auto => problem_logic_goals [map prop_of goal_cls]
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   631
			  | Fol => FOL
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   632
			  | Hol => HOL
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   633
	val keep_types = if is_fol_logic prob_logic then !fol_keep_types else is_typed_hol ()
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   634
	val classrel_clauses = if keep_types then ResClause.classrel_clauses_thy thy else []
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   635
	val arity_clauses = if keep_types then ResClause.arity_clause_thy thy else []
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   636
        val writer = if dfg then dfg_writer else tptp_writer 
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   637
	val file = atp_input_file()
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   638
    in
20131
c89ee2f4efd5 Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
mengj
parents: 20022
diff changeset
   639
	(writer prob_logic goal_cls file (axclauses,classrel_clauses,arity_clauses) user_lemmas_names;
19746
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   640
	 Output.debug ("Writing to " ^ file);
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   641
	 file)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   642
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   643
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   644
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   645
(**** remove tmp files ****)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   646
fun cond_rm_tmp file = 
19746
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   647
    if !keep_atp_input then Output.debug "ATP input kept..." 
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   648
    else if !destdir <> "" then Output.debug ("ATP input kept in directory " ^ (!destdir))
9ac97dc14214 warnings to debug outputs
paulson
parents: 19744
diff changeset
   649
    else (Output.debug "deleting ATP inputs..."; OS.FileSys.remove file);
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   650
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   651
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   652
(****** setup ATPs as Isabelle methods ******)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   653
fun atp_meth' tac ths ctxt = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   654
    Method.SIMPLE_METHOD' HEADGOAL
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   655
    (tac ctxt ths);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   656
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   657
fun atp_meth tac ths ctxt = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   658
    let val thy = ProofContext.theory_of ctxt
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   659
	val _ = ResClause.init thy
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   660
	val _ = ResHolClause.init thy
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   661
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   662
	atp_meth' tac ths ctxt
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   663
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   664
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   665
fun atp_method tac = Method.thms_ctxt_args (atp_meth tac);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   666
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   667
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   668
(* automatic ATP invocation                                    *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   669
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   670
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   671
(* call prover with settings and problem file for the current subgoal *)
17764
fde495b9e24b improved process handling. tidied
paulson
parents: 17746
diff changeset
   672
fun watcher_call_provers sign sg_terms (childin, childout, pid) =
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   673
  let
17422
3b237822985d massive tidy-up and simplification
paulson
parents: 17404
diff changeset
   674
    fun make_atp_list [] n = []
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   675
      | make_atp_list (sg_term::xs) n =
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   676
          let
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   677
            val probfile = prob_pathname n
17690
8ba7c3cd24a8 time limit option; fixed bug concerning first line of ATP output
paulson
parents: 17525
diff changeset
   678
            val time = Int.toString (!time_limit)
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   679
          in
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   680
            Output.debug ("problem file in watcher_call_provers is " ^ probfile);
17764
fde495b9e24b improved process handling. tidied
paulson
parents: 17746
diff changeset
   681
            (*options are separated by Watcher.setting_sep, currently #"%"*)
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   682
            if !prover = "spass"
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   683
            then
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   684
              let val spass = helper_path "SPASS_HOME" "SPASS"
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   685
                  val sopts =
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   686
   "-Auto%-SOS=1%-PGiven=0%-PProblem=0%-Splits=0%-FullRed=0%-DocProof%-TimeLimit=" ^ time
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   687
              in 
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   688
                  ("spass", spass, sopts, probfile) :: make_atp_list xs (n+1)
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   689
              end
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   690
            else if !prover = "vampire"
17235
8e55ad29b690 Added ECommunication.ML and modified res_atp.ML, Reconstruction.thy, and
quigley
parents: 17234
diff changeset
   691
	    then 
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
   692
              let val vampire = helper_path "VAMPIRE_HOME" "vampire"
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   693
                  val casc = if !time_limit > 70 then "--mode casc%" else ""
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   694
                  val vopts = casc ^ "-m 100000%-t " ^ time
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   695
              in
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   696
                  ("vampire", vampire, vopts, probfile) :: make_atp_list xs (n+1)
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   697
              end
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   698
      	     else if !prover = "E"
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   699
      	     then
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
   700
	       let val Eprover = helper_path "E_HOME" "eproof"
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   701
	       in
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   702
		  ("E", Eprover, 
19744
73aab222fecb Giving the "--silent" switch to E, to produce less output
paulson
parents: 19722
diff changeset
   703
		     "--tptp-in%-l5%-xAuto%-tAuto%--silent%--cpu-limit=" ^ time, probfile) ::
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   704
		   make_atp_list xs (n+1)
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   705
	       end
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   706
	     else error ("Invalid prover name: " ^ !prover)
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   707
          end
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   708
17422
3b237822985d massive tidy-up and simplification
paulson
parents: 17404
diff changeset
   709
    val atp_list = make_atp_list sg_terms 1
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   710
  in
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   711
    Watcher.callResProvers(childout,atp_list);
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   712
    Output.debug "Sent commands to watcher!"
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   713
  end
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   714
  
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   715
fun trace_array fname =
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   716
  let val path = File.tmp_path (Path.basic fname)
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   717
  in  Array.app (File.append path o (fn s => s ^ "\n"))  end;
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   718
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   719
(*We write out problem files for each subgoal. Argument probfile generates filenames,
18986
5060ca625e02 tidying
paulson
parents: 18863
diff changeset
   720
  and allows the suppression of the suffix "_1" in problem-generation mode.
5060ca625e02 tidying
paulson
parents: 18863
diff changeset
   721
  FIXME: does not cope with &&, and it isn't easy because one could have multiple
5060ca625e02 tidying
paulson
parents: 18863
diff changeset
   722
  subgoals, each involving &&.*)
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   723
fun write_problem_files probfile (ctxt,th)  =
18753
aa82bd41555d ResClasimp.get_clasimp_lemmas now takes all subgoals rather than only the first
paulson
parents: 18700
diff changeset
   724
  let val goals = Thm.prems_of th
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   725
      val _ = Output.debug ("number of subgoals = " ^ Int.toString (length goals))
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   726
      val (included_thms,white_thms) = get_clasimp_atp_lemmas ctxt []
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   727
      val rm_blacklist_cls = blacklist_filter included_thms
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   728
      val cla_simp_atp_clauses = ResAxioms.cnf_rules_pairs rm_blacklist_cls
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   729
      val axclauses = get_relevant_clauses ctxt cla_simp_atp_clauses (ResAxioms.cnf_rules_pairs white_thms) goals 
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   730
      val _ = Output.debug ("total clauses from thms = " ^ Int.toString (length axclauses))
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   731
      val thy = ProofContext.theory_of ctxt
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   732
      fun get_neg_subgoals n =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   733
	  if n=0 then []
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   734
	  else
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   735
	      let val st = Seq.hd (EVERY' [rtac ccontr, ObjectLogic.atomize_tac, 
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   736
	                                   skolemize_tac] n th)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   737
		  val negs = Option.valOf (metahyps_thms n st)
19442
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   738
		  val negs_clauses = make_clauses negs
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   739
	      in
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   740
		  negs_clauses :: get_neg_subgoals (n-1)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   741
	      end
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   742
      val neg_subgoals = get_neg_subgoals (length goals) 
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   743
      val goals_logic = case !linkup_logic_mode of
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   744
                            Auto => problem_logic_goals (map (map prop_of) neg_subgoals)
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   745
			  | Fol => FOL
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   746
			  | Hol => HOL
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   747
      val keep_types = if is_fol_logic goals_logic then !ResClause.keep_types else is_typed_hol ()
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   748
      val classrel_clauses = if keep_types then ResClause.classrel_clauses_thy thy else []
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   749
      val _ = Output.debug ("classrel clauses = " ^ Int.toString (length classrel_clauses))
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   750
      val arity_clauses = if keep_types then ResClause.arity_clause_thy thy else []
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   751
      val _ = Output.debug ("arity clauses = " ^ Int.toString (length arity_clauses))
19718
e709f643c578 Added support for DFG format, used by SPASS.
mengj
parents: 19675
diff changeset
   752
      val writer = if !prover = "spass" then dfg_writer else tptp_writer 
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   753
      fun write_all [] _ = []
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   754
	| write_all (sub::subgoals) k =
20131
c89ee2f4efd5 Pass user lemmas' names to ResHolClause.tptp_write_file and dfg_write_file.
mengj
parents: 20022
diff changeset
   755
	   (writer goals_logic sub (probfile k) (axclauses,classrel_clauses,arity_clauses) [],
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   756
	    probfile k) :: write_all subgoals (k-1)
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   757
      val (clnames::_, filenames) = ListPair.unzip (write_all neg_subgoals (length goals))
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   758
      val thm_names = Array.fromList clnames
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   759
      val _ = if !Output.show_debug_msgs 
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   760
              then trace_array "thm_names" thm_names else ()
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   761
  in
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   762
      (filenames, thm_names)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   763
  end;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   764
17775
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   765
val last_watcher_pid = ref (NONE : (TextIO.instream * TextIO.outstream * 
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   766
                                    Posix.Process.pid * string list) option);
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   767
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   768
fun kill_last_watcher () =
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   769
    (case !last_watcher_pid of 
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   770
         NONE => ()
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   771
       | SOME (_, _, pid, files) => 
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   772
	  (Output.debug ("Killing old watcher, pid = " ^ string_of_pid pid);
17775
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   773
	   Watcher.killWatcher pid;  
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   774
	   ignore (map (try OS.FileSys.remove) files)))
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   775
     handle OS.SysErr _ => Output.debug "Attempt to kill watcher failed";
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17502
diff changeset
   776
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17502
diff changeset
   777
(*writes out the current clasimpset to a tptp file;
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17502
diff changeset
   778
  turns off xsymbol at start of function, restoring it at end    *)
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
   779
val isar_atp = setmp print_mode [] 
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   780
 (fn (ctxt, th) =>
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   781
  if Thm.no_prems th then ()
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   782
  else
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   783
    let
17775
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   784
      val _ = kill_last_watcher()
19675
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   785
      val (files,thm_names) = write_problem_files prob_pathname (ctxt,th)
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   786
      val (childin, childout, pid) = Watcher.createWatcher (th, thm_names)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   787
    in
17772
818cec5f82a4 major simplification: removal of the goalstring argument
paulson
parents: 17764
diff changeset
   788
      last_watcher_pid := SOME (childin, childout, pid, files);
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   789
      Output.debug ("problem files: " ^ space_implode ", " files); 
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   790
      Output.debug ("pid: " ^ string_of_pid pid);
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   791
      watcher_call_provers (sign_of_thm th) (Thm.prems_of th) (childin, childout, pid)
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   792
    end);
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   793
17422
3b237822985d massive tidy-up and simplification
paulson
parents: 17404
diff changeset
   794
val isar_atp_writeonly = setmp print_mode [] 
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   795
      (fn (ctxt,th) =>
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   796
       if Thm.no_prems th then ()
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   797
       else 
20022
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   798
         let val probfile = if Thm.nprems_of th = 1 then probfile_nosuffix 
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   799
          	            else prob_pathname
b07a138b4e7d some tidying; fixed the output of theorem names
paulson
parents: 19894
diff changeset
   800
         in ignore (write_problem_files probfile (ctxt,th)) end);
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   801
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   802
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   803
(** the Isar toplevel hook **)
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   804
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   805
fun invoke_atp_ml (ctxt, goal) =
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   806
  let val thy = ProofContext.theory_of ctxt;
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   807
  in
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   808
    Output.debug ("subgoals in isar_atp:\n" ^ 
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   809
		  Pretty.string_of (ProofContext.pretty_term ctxt
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   810
		    (Logic.mk_conjunction_list (Thm.prems_of goal))));
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   811
    Output.debug ("current theory: " ^ Context.theory_name thy);
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 17091
diff changeset
   812
    hook_count := !hook_count +1;
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   813
    Output.debug ("in hook for time: " ^ Int.toString (!hook_count));
16925
0fd7b1438d28 simpler variable names, and no types for monomorphic constants
paulson
parents: 16904
diff changeset
   814
    ResClause.init thy;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   815
    ResHolClause.init thy;
17690
8ba7c3cd24a8 time limit option; fixed bug concerning first line of ATP output
paulson
parents: 17525
diff changeset
   816
    if !destdir = "" andalso !time_limit > 0 then isar_atp (ctxt, goal)
17502
8836793df947 further tidying; killing of old Watcher loops
paulson
parents: 17488
diff changeset
   817
    else isar_atp_writeonly (ctxt, goal)
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   818
  end;
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   819
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   820
val invoke_atp = Toplevel.no_timing o Toplevel.unknown_proof o Toplevel.keep
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   821
 (fn state =>
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   822
  let val (ctxt, (_, goal)) = Proof.get_goal (Toplevel.proof_of state)
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   823
  in  invoke_atp_ml (ctxt, goal)  end);
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   824
17091
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   825
val call_atpP =
17746
af59c748371d fixed the ascii-armouring of goalstring
paulson
parents: 17717
diff changeset
   826
  OuterSyntax.command 
17091
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   827
    "ProofGeneral.call_atp" 
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   828
    "call automatic theorem provers" 
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   829
    OuterKeyword.diag
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   830
    (Scan.succeed invoke_atp);
17091
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   831
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   832
val _ = OuterSyntax.add_parsers [call_atpP];
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   833
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   834
end;