src/HOL/Tools/res_atp.ML
author mengj
Thu, 25 May 2006 11:51:37 +0200
changeset 19722 e7a5b54248bc
parent 19718 e709f643c578
child 19744 73aab222fecb
permissions -rw-r--r--
Added in settings used by "spass" method.
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
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
    29
  val atp_method: (ProofContext.context -> thm list -> int -> Tactical.tactic) ->
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    30
		  Method.src -> ProofContext.context -> Method.method
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
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    39
  val run_relevance_filter: bool ref
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
    40
  val invoke_atp_ml : ProofContext.context * thm -> unit
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    41
  val add_claset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    42
  val add_simpset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    43
  val add_clasimp : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    44
  val add_atpset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    45
  val rm_claset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    46
  val rm_simpset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    47
  val rm_atpset : unit -> unit
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
    48
  val rm_clasimp : unit -> unit
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    49
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    50
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
    51
structure ResAtp : RES_ATP =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    52
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    53
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    54
(********************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    55
(* some settings for both background automatic ATP calling procedure*)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    56
(* and also explicit ATP invocation methods                         *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    57
(********************************************************************)
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
(*** background linkup ***)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    60
val call_atp = ref false; 
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 17091
diff changeset
    61
val hook_count = ref 0;
18675
333a73034023 more practical time limit
paulson
parents: 18404
diff changeset
    62
val time_limit = ref 30;
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17317
diff changeset
    63
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
    64
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
    65
      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
    66
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
    67
val problem_name = ref "prob";
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    68
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    69
(*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
    70
  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
    71
fun helper_path evar base =
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    72
  case getenv evar of
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    73
      "" => error  ("Isabelle environment variable " ^ evar ^ " not defined")
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    74
    | home => 
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    75
        let val path = home ^ "/" ^ base
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    76
        in  if File.exists (File.unpack_platform_path path) then path 
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    77
	    else error ("Could not find the file " ^ path)
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    78
	end;  
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
    79
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
    80
fun probfile_nosuffix _ = 
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    81
  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
    82
  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
    83
  then !destdir ^ "/" ^ !problem_name
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17435
diff changeset
    84
  else error ("No such directory: " ^ !destdir);
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    85
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
    86
fun prob_pathname n = probfile_nosuffix n ^ "_" ^ Int.toString n;
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
    87
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    88
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    89
(*** ATP methods ***)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    90
val vampire_time = ref 60;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    91
val eprover_time = ref 60;
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    92
val spass_time = ref 60;
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
    93
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    94
fun run_vampire time =  
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    95
    if (time >0) then vampire_time:= time
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    96
    else vampire_time:=60;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    97
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    98
fun run_eprover time = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    99
    if (time > 0) then eprover_time:= time
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   100
    else eprover_time:=60;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   101
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   102
fun run_spass time = 
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   103
    if (time > 0) then spass_time:=time
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   104
    else spass_time:=60;
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   105
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   106
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   107
fun vampireLimit () = !vampire_time;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   108
fun eproverLimit () = !eprover_time;
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   109
fun spassLimit () = !spass_time;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   110
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   111
val keep_atp_input = ref false;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   112
val fol_keep_types = ResClause.keep_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   113
val hol_full_types = ResHolClause.full_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   114
val hol_partial_types = ResHolClause.partial_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   115
val hol_const_types_only = ResHolClause.const_types_only;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   116
val hol_no_types = ResHolClause.no_types;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   117
fun hol_typ_level () = ResHolClause.find_typ_level ();
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   118
fun is_typed_hol () = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   119
    let val tp_level = hol_typ_level()
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   120
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   121
	not (tp_level = ResHolClause.T_NONE)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   122
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   123
val include_combS = ResHolClause.include_combS;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   124
val include_min_comb = ResHolClause.include_min_comb;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   125
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   126
fun atp_input_file () =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   127
    let val file = !problem_name 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   128
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   129
	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
   130
	else if File.exists (File.unpack_platform_path (!destdir))
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   131
	then !destdir ^ "/" ^ file
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   132
	else error ("No such directory: " ^ !destdir)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   133
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   134
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   135
val include_simpset = ref false;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   136
val include_claset = ref false; 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   137
val include_atpset = ref true;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   138
val add_simpset = (fn () => include_simpset:=true);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   139
val add_claset = (fn () => include_claset:=true);
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   140
val add_clasimp = (fn () => (include_simpset:=true;include_claset:=true));
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   141
val add_atpset = (fn () => include_atpset:=true);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   142
val rm_simpset = (fn () => include_simpset:=false);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   143
val rm_claset = (fn () => include_claset:=false);
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   144
val rm_clasimp = (fn () => (include_simpset:=false;include_claset:=false));
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   145
val rm_atpset = (fn () => include_atpset:=false);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   146
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   147
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   148
(**** relevance filter ****)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   149
val run_relevance_filter = ref true;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   150
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   151
(******************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   152
(* 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
   153
(******************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   154
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   155
datatype logic = FOL | HOL | HOLC | HOLCS;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   156
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   157
fun string_of_logic FOL = "FOL"
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   158
  | string_of_logic HOL = "HOL"
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   159
  | string_of_logic HOLC = "HOLC"
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   160
  | string_of_logic HOLCS = "HOLCS";
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   161
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   162
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   163
fun is_fol_logic FOL = true
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   164
  | is_fol_logic  _ = false
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
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   167
(*HOLCS will not occur here*)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   168
fun upgrade_lg HOLC _ = HOLC
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   169
  | upgrade_lg HOL HOLC = HOLC
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   170
  | upgrade_lg HOL _ = HOL
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   171
  | upgrade_lg FOL lg = lg; 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   172
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   173
(* check types *)
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   174
fun has_bool_hfn (Type("bool",_)) = true
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   175
  | has_bool_hfn (Type("fun",_)) = true
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   176
  | has_bool_hfn (Type(_, Ts)) = exists has_bool_hfn Ts
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   177
  | has_bool_hfn _ = false;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   178
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   179
fun is_hol_fn tp =
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   180
    let val (targs,tr) = strip_type tp
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   181
    in
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   182
	exists (has_bool_hfn) (tr::targs)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   183
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   184
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   185
fun is_hol_pred tp =
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   186
    let val (targs,tr) = strip_type tp
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   187
    in
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   188
	exists (has_bool_hfn) targs
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   189
    end;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   190
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   191
exception FN_LG of term;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   192
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   193
fun fn_lg (t as Const(f,tp)) (lg,seen) = 
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   194
    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
   195
  | fn_lg (t as Free(f,tp)) (lg,seen) = 
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   196
    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
   197
  | fn_lg (t as Var(f,tp)) (lg,seen) =
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   198
    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
   199
  | 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
   200
  | fn_lg f _ = raise FN_LG(f); 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   201
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   202
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   203
fun term_lg [] (lg,seen) = (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   204
  | term_lg (tm::tms) (FOL,seen) =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   205
    let val (f,args) = strip_comb tm
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   206
	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
   207
			  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
   208
	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
   209
          if is_fol_logic lg' then ()
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
   210
          else warning ("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
   211
	 in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   212
	     term_lg (args@tms) (lg',seen')
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   213
    end
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   214
  | term_lg _ (lg,seen) = (lg,seen)
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
exception PRED_LG of term;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   217
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   218
fun pred_lg (t as Const(P,tp)) (lg,seen)= 
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   219
    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
   220
  | pred_lg (t as Free(P,tp)) (lg,seen) =
19451
c7a25c05986c Changed the logic detection method.
mengj
parents: 19450
diff changeset
   221
    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
   222
  | 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
   223
  | pred_lg P _ = raise PRED_LG(P);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   224
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   225
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   226
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
   227
  | lit_lg P (lg,seen) =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   228
    let val (pred,args) = strip_comb P
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   229
	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
   230
			  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
   231
	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
   232
          if is_fol_logic lg' then ()
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
   233
          else warning ("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
   234
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   235
	term_lg args (lg',seen')
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   236
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   237
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   238
fun lits_lg [] (lg,seen) = (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   239
  | lits_lg (lit::lits) (FOL,seen) =
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   240
    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
   241
	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
   242
          if is_fol_logic lg then ()
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
   243
          else warning ("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
   244
    in
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   245
	lits_lg lits (lg,seen')
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   246
    end
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   247
  | lits_lg lits (lg,seen) = (lg,seen);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   248
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   249
19227
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   250
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
   251
    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
   252
  | dest_disj_aux t disjs = t::disjs;
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   253
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   254
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
   255
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   256
fun logic_of_clause tm (lg,seen) =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   257
    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
   258
	val disjs = dest_disj tm'
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   259
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   260
	lits_lg disjs (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   261
    end;
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
fun logic_of_clauses [] (lg,seen) = (lg,seen)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   264
  | 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
   265
    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
   266
	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
   267
          if is_fol_logic lg then ()
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
   268
          else warning ("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
   269
    in
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   270
	logic_of_clauses clss (lg,seen')
d15f2baa7ecc Added more functions to the signature and tidied up some functions.
mengj
parents: 19205
diff changeset
   271
    end
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   272
  | logic_of_clauses (cls::clss) (lg,seen) = (lg,seen);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   273
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   274
fun problem_logic_goals_aux [] (lg,seen) = lg
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   275
  | problem_logic_goals_aux (subgoal::subgoals) (lg,seen) = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   276
    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
   277
    
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   278
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
   279
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   280
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   281
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   282
(* ATP invocation methods setup                                *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   283
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   284
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   285
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   286
(**** prover-specific format: TPTP ****)
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
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   289
fun cnf_hyps_thms ctxt = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   290
    let val ths = ProofContext.prems_of ctxt
19617
7cb4b67d4b97 avoid raw equality on type thm;
wenzelm
parents: 19490
diff changeset
   291
    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
   292
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   293
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   294
(**** write to files ****)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   295
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   296
datatype mode = Auto | Fol | Hol;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   297
19450
651d949776f8 exported linkup_logic_mode and changed the default setting
paulson
parents: 19445
diff changeset
   298
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
   299
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   300
fun tptp_writer logic goals filename (axioms,classrels,arities) =
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   301
    if is_fol_logic logic 
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   302
    then ResClause.tptp_write_file goals filename (axioms, classrels, arities)
19490
bf7f8347174a removed the functions for getting HOL helper paths.
mengj
parents: 19451
diff changeset
   303
    else ResHolClause.tptp_write_file goals filename (axioms, classrels, arities);
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   304
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   305
fun dfg_writer logic goals filename (axioms,classrels,arities) =
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   306
    if is_fol_logic logic 
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   307
    then ResClause.dfg_write_file goals filename (axioms, classrels, arities)
19718
e709f643c578 Added support for DFG format, used by SPASS.
mengj
parents: 19675
diff changeset
   308
    else ResHolClause.dfg_write_file goals filename (axioms, classrels, arities);
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   309
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   310
19722
e7a5b54248bc Added in settings used by "spass" method.
mengj
parents: 19718
diff changeset
   311
fun write_subgoal_file dfg mode ctxt conjectures user_thms n =
19442
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   312
    let val conj_cls = make_clauses conjectures 
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   313
	val hyp_cls = cnf_hyps_thms ctxt
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   314
	val goal_cls = conj_cls@hyp_cls
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   315
	val user_rules = map ResAxioms.pairname user_thms
19675
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   316
	val axclauses_as_thms = ResClasimp.get_clasimp_atp_lemmas ctxt (map prop_of goal_cls) user_rules (!include_claset,!include_simpset,!include_atpset) (!run_relevance_filter)  
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   317
	val thy = ProofContext.theory_of ctxt
19442
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   318
	val prob_logic = case mode of Auto => problem_logic_goals [map prop_of goal_cls]
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   319
				    | Fol => FOL
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   320
				    | Hol => HOL
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   321
	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
   322
	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
   323
	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
   324
        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
   325
	val file = atp_input_file()
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   326
    in
19352
1a07f6cf1e6c lemmas returned from ResClasimp.get_clasimp_atp_lemmas are thm rather than term.
mengj
parents: 19227
diff changeset
   327
	(writer prob_logic goal_cls file (axclauses_as_thms,classrel_clauses,arity_clauses);
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   328
	 warning ("Writing to " ^ file);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   329
	 file)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   330
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   331
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   332
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   333
(**** remove tmp files ****)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   334
fun cond_rm_tmp file = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   335
    if !keep_atp_input then warning "ATP input kept..." 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   336
    else if !destdir <> "" then warning ("ATP input kept in directory " ^ (!destdir))
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   337
    else (warning "deleting ATP inputs..."; OS.FileSys.remove file);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   338
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   339
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   340
(****** setup ATPs as Isabelle methods ******)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   341
fun atp_meth' tac ths ctxt = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   342
    Method.SIMPLE_METHOD' HEADGOAL
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   343
    (tac ctxt ths);
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   344
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   345
fun atp_meth tac ths ctxt = 
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   346
    let val thy = ProofContext.theory_of ctxt
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   347
	val _ = ResClause.init thy
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   348
	val _ = ResHolClause.init thy
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   349
    in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   350
	atp_meth' tac ths ctxt
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   351
    end;
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   352
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   353
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
   354
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   355
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   356
(* automatic ATP invocation                                    *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   357
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   358
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   359
(* call prover with settings and problem file for the current subgoal *)
17764
fde495b9e24b improved process handling. tidied
paulson
parents: 17746
diff changeset
   360
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
   361
  let
17422
3b237822985d massive tidy-up and simplification
paulson
parents: 17404
diff changeset
   362
    fun make_atp_list [] n = []
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   363
      | 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
   364
          let
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   365
            val probfile = prob_pathname n
17690
8ba7c3cd24a8 time limit option; fixed bug concerning first line of ATP output
paulson
parents: 17525
diff changeset
   366
            val time = Int.toString (!time_limit)
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   367
          in
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   368
            Output.debug ("problem file in watcher_call_provers is " ^ probfile);
17764
fde495b9e24b improved process handling. tidied
paulson
parents: 17746
diff changeset
   369
            (*options are separated by Watcher.setting_sep, currently #"%"*)
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   370
            if !prover = "spass"
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   371
            then
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   372
              let val spass = helper_path "SPASS_HOME" "SPASS"
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   373
                  val sopts =
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   374
   "-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
   375
              in 
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   376
                  ("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
   377
              end
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   378
            else if !prover = "vampire"
17235
8e55ad29b690 Added ECommunication.ML and modified res_atp.ML, Reconstruction.thy, and
quigley
parents: 17234
diff changeset
   379
	    then 
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
   380
              let val vampire = helper_path "VAMPIRE_HOME" "vampire"
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   381
                  val casc = if !time_limit > 70 then "--mode casc%" else ""
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   382
                  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
   383
              in
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   384
                  ("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
   385
              end
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   386
      	     else if !prover = "E"
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   387
      	     then
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17775
diff changeset
   388
	       let val Eprover = helper_path "E_HOME" "eproof"
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   389
	       in
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   390
		  ("E", Eprover, 
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   391
		     "--tptp-in%-l5%-xAuto%-tAuto%--cpu-limit=" ^ time, probfile) ::
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   392
		   make_atp_list xs (n+1)
17306
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   393
	       end
5cde710a8a23 Progress on eprover linkup, also massive tidying
paulson
parents: 17305
diff changeset
   394
	     else error ("Invalid prover name: " ^ !prover)
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   395
          end
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   396
17422
3b237822985d massive tidy-up and simplification
paulson
parents: 17404
diff changeset
   397
    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
   398
  in
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   399
    Watcher.callResProvers(childout,atp_list);
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   400
    Output.debug "Sent commands to watcher!"
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   401
  end
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   402
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   403
(*We write out problem files for each subgoal. Argument pf generates filenames,
18986
5060ca625e02 tidying
paulson
parents: 18863
diff changeset
   404
  and allows the suppression of the suffix "_1" in problem-generation mode.
5060ca625e02 tidying
paulson
parents: 18863
diff changeset
   405
  FIXME: does not cope with &&, and it isn't easy because one could have multiple
5060ca625e02 tidying
paulson
parents: 18863
diff changeset
   406
  subgoals, each involving &&.*)
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   407
fun write_problem_files pf (ctxt,th)  =
18753
aa82bd41555d ResClasimp.get_clasimp_lemmas now takes all subgoals rather than only the first
paulson
parents: 18700
diff changeset
   408
  let val goals = Thm.prems_of th
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   409
      val _ = Output.debug ("number of subgoals = " ^ Int.toString (length goals))
19675
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   410
      val axclauses = ResClasimp.get_clasimp_atp_lemmas ctxt goals [] (true,true,true) (!run_relevance_filter) (* no user supplied rules here, because no user invocation *)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   411
      val _ = Output.debug ("claset, simprules and atprules total clauses = " ^ 
19675
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   412
                     Int.toString (length axclauses))
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   413
      val thy = ProofContext.theory_of ctxt
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   414
      fun get_neg_subgoals n =
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   415
	  if n=0 then []
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   416
	  else
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   417
	      let val st = Seq.hd (EVERY'
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   418
				       [rtac ccontr, ObjectLogic.atomize_tac, skolemize_tac] n th)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   419
		  val negs = Option.valOf (metahyps_thms n st)
19442
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   420
		  val negs_clauses = make_clauses negs
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   421
	      in
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   422
		  negs_clauses::(get_neg_subgoals (n - 1))
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   423
	      end
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   424
      val neg_subgoals = get_neg_subgoals (length goals) 
19442
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   425
      val goals_logic = case !linkup_logic_mode of Auto => problem_logic_goals (map (map prop_of) neg_subgoals)
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   426
						 | Fol => FOL
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   427
						 | Hol => HOL
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   428
      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
   429
      val classrel_clauses = if keep_types then ResClause.classrel_clauses_thy thy else []
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   430
      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
   431
      val arity_clauses = if keep_types then ResClause.arity_clause_thy thy else []
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   432
      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
   433
      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
   434
      fun write_all [] _ = []
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   435
	| write_all (subgoal::subgoals) k =
19442
ad8bb8346e51 Tidied up some programs.
mengj
parents: 19352
diff changeset
   436
	  (writer goals_logic subgoal (pf k) (axclauses,classrel_clauses,arity_clauses); pf k):: (write_all subgoals (k - 1))
19675
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   437
      val thm_names = Array.fromList (map (#1 o #2) axclauses)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   438
  in
19675
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   439
      (write_all neg_subgoals (length goals), thm_names)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   440
  end;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   441
17775
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   442
val last_watcher_pid = ref (NONE : (TextIO.instream * TextIO.outstream * 
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   443
                                    Posix.Process.pid * string list) option);
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   444
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   445
fun kill_last_watcher () =
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   446
    (case !last_watcher_pid of 
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   447
         NONE => ()
19445
da75577642a9 tidying; ATP options including CASC mode for Vampire
paulson
parents: 19442
diff changeset
   448
       | SOME (_, _, pid, files) => 
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   449
	  (Output.debug ("Killing old watcher, pid = " ^ string_of_pid pid);
17775
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   450
	   Watcher.killWatcher pid;  
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   451
	   ignore (map (try OS.FileSys.remove) files)))
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   452
     handle OS.SysErr _ => Output.debug "Attempt to kill watcher failed";
17525
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17502
diff changeset
   453
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17502
diff changeset
   454
(*writes out the current clasimpset to a tptp file;
ae5bb6001afb tidying, and support for axclass/classrel clauses
paulson
parents: 17502
diff changeset
   455
  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
   456
val isar_atp = setmp print_mode [] 
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   457
 (fn (ctxt, th) =>
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   458
  if Thm.no_prems th then ()
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   459
  else
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   460
    let
17775
2679ba74411f minor code tidyig
paulson
parents: 17773
diff changeset
   461
      val _ = kill_last_watcher()
19675
a4894fb2a5f2 removing the string array from the result of get_clasimp_atp_lemmas
paulson
parents: 19641
diff changeset
   462
      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
   463
      val (childin, childout, pid) = Watcher.createWatcher (th, thm_names)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   464
    in
17772
818cec5f82a4 major simplification: removal of the goalstring argument
paulson
parents: 17764
diff changeset
   465
      last_watcher_pid := SOME (childin, childout, pid, files);
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   466
      Output.debug ("problem files: " ^ space_implode ", " files); 
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   467
      Output.debug ("pid: " ^ string_of_pid pid);
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   468
      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
   469
    end);
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   470
17422
3b237822985d massive tidy-up and simplification
paulson
parents: 17404
diff changeset
   471
val isar_atp_writeonly = setmp print_mode [] 
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   472
      (fn (ctxt,th) =>
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   473
       if Thm.no_prems th then ()
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   474
       else 
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   475
         let val pf = if Thm.nprems_of th = 1 then probfile_nosuffix 
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   476
         	      else prob_pathname
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   477
         in ignore (write_problem_files pf (ctxt,th)) end);
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   478
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   479
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   480
(** the Isar toplevel hook **)
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
   481
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   482
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
   483
  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
   484
  in
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   485
    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
   486
		  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
   487
		    (Logic.mk_conjunction_list (Thm.prems_of goal))));
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   488
    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
   489
    hook_count := !hook_count +1;
18680
677e2bdd75f0 Output.debug;
wenzelm
parents: 18675
diff changeset
   490
    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
   491
    ResClause.init thy;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   492
    ResHolClause.init thy;
17690
8ba7c3cd24a8 time limit option; fixed bug concerning first line of ATP output
paulson
parents: 17525
diff changeset
   493
    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
   494
    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
   495
  end;
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   496
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   497
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
   498
 (fn state =>
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   499
  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
   500
  in  invoke_atp_ml (ctxt, goal)  end);
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   501
17091
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   502
val call_atpP =
17746
af59c748371d fixed the ascii-armouring of goalstring
paulson
parents: 17717
diff changeset
   503
  OuterSyntax.command 
17091
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   504
    "ProofGeneral.call_atp" 
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   505
    "call automatic theorem provers" 
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   506
    OuterKeyword.diag
19205
4ec788c69f82 Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents: 19194
diff changeset
   507
    (Scan.succeed invoke_atp);
17091
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   508
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   509
val _ = OuterSyntax.add_parsers [call_atpP];
13593aa6a546 new command to invoke ATPs
paulson
parents: 16955
diff changeset
   510
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   511
end;