| author | blanchet | 
| Mon, 21 May 2012 10:39:32 +0200 | |
| changeset 47946 | 33afcfad3f8d | 
| parent 47912 | 12de57c5eee5 | 
| child 48130 | defbcdc60fd6 | 
| permissions | -rw-r--r-- | 
| 46320 | 1  | 
(* Title: HOL/Tools/Metis/metis_generate.ML  | 
| 38027 | 2  | 
Author: Jia Meng, Cambridge University Computer Laboratory and NICTA  | 
| 
39497
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
3  | 
Author: Kong W. Susanto, Cambridge University Computer Laboratory  | 
| 
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
4  | 
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory  | 
| 
36393
 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 
blanchet 
parents: 
36378 
diff
changeset
 | 
5  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 15347 | 6  | 
|
| 
39494
 
bf7dd4902321
rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
 
blanchet 
parents: 
39355 
diff
changeset
 | 
7  | 
Translation of HOL to FOL for Metis.  | 
| 15347 | 8  | 
*)  | 
9  | 
||
| 46320 | 10  | 
signature METIS_GENERATE =  | 
| 24310 | 11  | 
sig  | 
| 46320 | 12  | 
type type_enc = ATP_Problem_Generate.type_enc  | 
| 
44411
 
e3629929b171
change Metis's default settings if type information axioms are generated
 
blanchet 
parents: 
44409 
diff
changeset
 | 
13  | 
|
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
14  | 
datatype isa_thm =  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
15  | 
Isa_Reflexive_or_Trivial |  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
16  | 
Isa_Lambda_Lifted |  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
17  | 
Isa_Raw of thm  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
18  | 
|
| 43094 | 19  | 
val metis_equal : string  | 
20  | 
val metis_predicator : string  | 
|
21  | 
val metis_app_op : string  | 
|
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
22  | 
val metis_systematic_type_tag : string  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
23  | 
val metis_ad_hoc_type_tag : string  | 
| 
42098
 
f978caf60bbe
more robust handling of variables in new Skolemizer
 
blanchet 
parents: 
41491 
diff
changeset
 | 
24  | 
val metis_generated_var_prefix : string  | 
| 
43231
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
25  | 
val trace : bool Config.T  | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
26  | 
val verbose : bool Config.T  | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
27  | 
val trace_msg : Proof.context -> (unit -> string) -> unit  | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
28  | 
val verbose_warning : Proof.context -> string -> unit  | 
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
29  | 
val metis_name_table : ((string * int) * ((type_enc -> string) * bool)) list  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
30  | 
val reveal_old_skolem_terms : (string * term) list -> term -> term  | 
| 
45569
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
31  | 
val reveal_lam_lifted : (string * term) list -> term -> term  | 
| 40157 | 32  | 
val prepare_metis_problem :  | 
| 45508 | 33  | 
Proof.context -> type_enc -> string -> thm list -> thm list  | 
34  | 
-> int Symtab.table * (Metis_Thm.thm * isa_thm) list  | 
|
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
46409 
diff
changeset
 | 
35  | 
* (unit -> (string * int) list)  | 
| 45508 | 36  | 
* ((string * term) list * (string * term) list)  | 
| 24310 | 37  | 
end  | 
| 15347 | 38  | 
|
| 46320 | 39  | 
structure Metis_Generate : METIS_GENERATE =  | 
| 15347 | 40  | 
struct  | 
41  | 
||
| 
43092
 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 
blanchet 
parents: 
43091 
diff
changeset
 | 
42  | 
open ATP_Problem  | 
| 46320 | 43  | 
open ATP_Problem_Generate  | 
| 15347 | 44  | 
|
| 43094 | 45  | 
val metis_equal = "="  | 
46  | 
val metis_predicator = "{}"
 | 
|
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
47  | 
val metis_app_op = Metis_Name.toString Metis_Term.appName  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
48  | 
val metis_systematic_type_tag =  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
49  | 
Metis_Name.toString Metis_Term.hasTypeFunctionName  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
50  | 
val metis_ad_hoc_type_tag = "**"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43003 
diff
changeset
 | 
51  | 
val metis_generated_var_prefix = "_"  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
52  | 
|
| 
43231
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
53  | 
val trace = Attrib.setup_config_bool @{binding metis_trace} (K false)
 | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
54  | 
val verbose = Attrib.setup_config_bool @{binding metis_verbose} (K true)
 | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
55  | 
|
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
56  | 
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()  | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
57  | 
fun verbose_warning ctxt msg =  | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
58  | 
  if Config.get ctxt verbose then warning ("Metis: " ^ msg) else ()
 | 
| 
 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 
blanchet 
parents: 
43224 
diff
changeset
 | 
59  | 
|
| 43094 | 60  | 
val metis_name_table =  | 
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
61  | 
[((tptp_equal, 2), (K metis_equal, false)),  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
62  | 
((tptp_old_equal, 2), (K metis_equal, false)),  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
63  | 
((prefixed_predicator_name, 1), (K metis_predicator, false)),  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
64  | 
((prefixed_app_op_name, 2), (K metis_app_op, false)),  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
65  | 
((prefixed_type_tag_name, 2),  | 
| 44782 | 66  | 
(fn type_enc =>  | 
67  | 
if level_of_type_enc type_enc = All_Types then metis_systematic_type_tag  | 
|
68  | 
else metis_ad_hoc_type_tag, true))]  | 
|
| 43094 | 69  | 
|
| 
39896
 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 
blanchet 
parents: 
39890 
diff
changeset
 | 
70  | 
fun old_skolem_const_name i j num_T_args =  | 
| 
 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 
blanchet 
parents: 
39890 
diff
changeset
 | 
71  | 
old_skolem_const_prefix ^ Long_Name.separator ^  | 
| 41491 | 72  | 
(space_implode Long_Name.separator (map string_of_int [i, j, num_T_args]))  | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
73  | 
|
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
74  | 
fun conceal_old_skolem_terms i old_skolems t =  | 
| 
39953
 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 
blanchet 
parents: 
39946 
diff
changeset
 | 
75  | 
  if exists_Const (curry (op =) @{const_name Meson.skolem} o fst) t then
 | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
76  | 
let  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
77  | 
fun aux old_skolems  | 
| 
39953
 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 
blanchet 
parents: 
39946 
diff
changeset
 | 
78  | 
             (t as (Const (@{const_name Meson.skolem}, Type (_, [_, T])) $ _)) =
 | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
79  | 
let  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
80  | 
val (old_skolems, s) =  | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
81  | 
if i = ~1 then  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
82  | 
                (old_skolems, @{const_name undefined})
 | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
83  | 
else case AList.find (op aconv) old_skolems t of  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
84  | 
s :: _ => (old_skolems, s)  | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
85  | 
| [] =>  | 
| 
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
86  | 
let  | 
| 
39896
 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 
blanchet 
parents: 
39890 
diff
changeset
 | 
87  | 
val s = old_skolem_const_name i (length old_skolems)  | 
| 
 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 
blanchet 
parents: 
39890 
diff
changeset
 | 
88  | 
(length (Term.add_tvarsT T []))  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
89  | 
in ((s, t) :: old_skolems, s) end  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
90  | 
in (old_skolems, Const (s, T)) end  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
91  | 
| aux old_skolems (t1 $ t2) =  | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
92  | 
let  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
93  | 
val (old_skolems, t1) = aux old_skolems t1  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
94  | 
val (old_skolems, t2) = aux old_skolems t2  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
95  | 
in (old_skolems, t1 $ t2) end  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
96  | 
| aux old_skolems (Abs (s, T, t')) =  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
97  | 
let val (old_skolems, t') = aux old_skolems t' in  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
98  | 
(old_skolems, Abs (s, T, t'))  | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
99  | 
end  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
100  | 
| aux old_skolems t = (old_skolems, t)  | 
| 
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
101  | 
in aux old_skolems t end  | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
102  | 
else  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
103  | 
(old_skolems, t)  | 
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
104  | 
|
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
105  | 
fun reveal_old_skolem_terms old_skolems =  | 
| 37632 | 106  | 
map_aterms (fn t as Const (s, _) =>  | 
| 
39896
 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 
blanchet 
parents: 
39890 
diff
changeset
 | 
107  | 
if String.isPrefix old_skolem_const_prefix s then  | 
| 
39886
 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 
blanchet 
parents: 
39720 
diff
changeset
 | 
108  | 
AList.lookup (op =) old_skolems s |> the  | 
| 
43826
 
2b094d17f432
fix subtle type inference bug in Metis -- different occurrences of the same Skolem might need to be typed differently, using paramify_vars overconstraints the typing problem
 
blanchet 
parents: 
43626 
diff
changeset
 | 
109  | 
|> map_types (map_type_tvar (K dummyT))  | 
| 37632 | 110  | 
else  | 
111  | 
t  | 
|
112  | 
| t => t)  | 
|
113  | 
||
| 
45569
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
114  | 
fun reveal_lam_lifted lambdas =  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
115  | 
map_aterms (fn t as Const (s, _) =>  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
116  | 
if String.isPrefix lam_lifted_prefix s then  | 
| 45508 | 117  | 
case AList.lookup (op =) lambdas s of  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
118  | 
SOME t =>  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
119  | 
                     Const (@{const_name Metis.lambda}, dummyT)
 | 
| 
45565
 
9c335d69a362
fixed bugs in lambda-lifting code -- ensure distinct names for variables
 
blanchet 
parents: 
45554 
diff
changeset
 | 
120  | 
$ map_types (map_type_tvar (K dummyT))  | 
| 
45569
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
121  | 
(reveal_lam_lifted lambdas t)  | 
| 45508 | 122  | 
| NONE => t  | 
123  | 
else  | 
|
124  | 
t  | 
|
125  | 
| t => t)  | 
|
126  | 
||
| 
37577
 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 
blanchet 
parents: 
37575 
diff
changeset
 | 
127  | 
|
| 
39497
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
128  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
129  | 
(* Logic maps manage the interface between HOL and first-order logic. *)  | 
| 
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
130  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
131  | 
|
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
132  | 
datatype isa_thm =  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
133  | 
Isa_Reflexive_or_Trivial |  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
134  | 
Isa_Lambda_Lifted |  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
135  | 
Isa_Raw of thm  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
136  | 
|
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
137  | 
val proxy_defs = map (fst o snd o snd) proxy_table  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
138  | 
val prepare_helper =  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
139  | 
Meson.make_meta_clause #> rewrite_rule (map safe_mk_meta_eq proxy_defs)  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43132 
diff
changeset
 | 
140  | 
|
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
141  | 
fun metis_term_from_atp type_enc (ATerm (s, tms)) =  | 
| 43094 | 142  | 
if is_tptp_variable s then  | 
| 
43268
 
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
 
blanchet 
parents: 
43259 
diff
changeset
 | 
143  | 
Metis_Term.Var (Metis_Name.fromString s)  | 
| 43094 | 144  | 
else  | 
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
145  | 
(case AList.lookup (op =) metis_name_table (s, length tms) of  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
146  | 
SOME (f, swap) => (f type_enc, swap)  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
147  | 
| NONE => (s, false))  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
148  | 
|> (fn (s, swap) =>  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
149  | 
Metis_Term.Fn (Metis_Name.fromString s,  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
150  | 
tms |> map (metis_term_from_atp type_enc)  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
151  | 
|> swap ? rev))  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
152  | 
fun metis_atom_from_atp type_enc (AAtom tm) =  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
153  | 
(case metis_term_from_atp type_enc tm of  | 
| 
43104
 
81d1b15aa0ae
use ":" for type information (looks good in Metis's output) and handle it in new path finder
 
blanchet 
parents: 
43103 
diff
changeset
 | 
154  | 
Metis_Term.Fn x => x  | 
| 
 
81d1b15aa0ae
use ":" for type information (looks good in Metis's output) and handle it in new path finder
 
blanchet 
parents: 
43103 
diff
changeset
 | 
155  | 
| _ => raise Fail "non CNF -- expected function")  | 
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
156  | 
| metis_atom_from_atp _ _ = raise Fail "not CNF -- expected atom"  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
157  | 
fun metis_literal_from_atp type_enc (AConn (ANot, [phi])) =  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
158  | 
(false, metis_atom_from_atp type_enc phi)  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
159  | 
| metis_literal_from_atp type_enc phi =  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
160  | 
(true, metis_atom_from_atp type_enc phi)  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
161  | 
fun metis_literals_from_atp type_enc (AConn (AOr, phis)) =  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
162  | 
maps (metis_literals_from_atp type_enc) phis  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
163  | 
| metis_literals_from_atp type_enc phi = [metis_literal_from_atp type_enc phi]  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
164  | 
fun metis_axiom_from_atp type_enc clauses (Formula (ident, _, phi, _, _)) =  | 
| 43173 | 165  | 
let  | 
166  | 
fun some isa =  | 
|
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
167  | 
SOME (phi |> metis_literals_from_atp type_enc  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
168  | 
|> Metis_LiteralSet.fromList  | 
| 43173 | 169  | 
|> Metis_Thm.axiom, isa)  | 
170  | 
in  | 
|
| 
47046
 
62ca06cc5a99
remove two options that were found to play hardly any role
 
blanchet 
parents: 
47039 
diff
changeset
 | 
171  | 
if String.isPrefix tags_sym_formula_prefix ident then  | 
| 43173 | 172  | 
Isa_Reflexive_or_Trivial |> some  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
173  | 
else if String.isPrefix conjecture_prefix ident then  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
174  | 
NONE  | 
| 43173 | 175  | 
else if String.isPrefix helper_prefix ident then  | 
| 43194 | 176  | 
case (String.isSuffix typed_helper_suffix ident,  | 
177  | 
space_explode "_" ident) of  | 
|
178  | 
(needs_fairly_sound, _ :: const :: j :: _) =>  | 
|
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47912 
diff
changeset
 | 
179  | 
nth (AList.lookup (op =) helper_table (const, needs_fairly_sound)  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47912 
diff
changeset
 | 
180  | 
|> the)  | 
| 43173 | 181  | 
(the (Int.fromString j) - 1)  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47912 
diff
changeset
 | 
182  | 
|> snd |> prepare_helper  | 
| 43194 | 183  | 
|> Isa_Raw |> some  | 
| 43173 | 184  | 
        | _ => raise Fail ("malformed helper identifier " ^ quote ident)
 | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
185  | 
else case try (unprefix fact_prefix) ident of  | 
| 43173 | 186  | 
SOME s =>  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
187  | 
let val s = s |> space_explode "_" |> tl |> space_implode "_"  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
188  | 
in  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
189  | 
case Int.fromString s of  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
190  | 
SOME j =>  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
191  | 
Meson.make_meta_clause (snd (nth clauses j)) |> Isa_Raw |> some  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
192  | 
| NONE =>  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
193  | 
if String.isPrefix lam_fact_prefix (unascii_of s) then  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
194  | 
Isa_Lambda_Lifted |> some  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
195  | 
else  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
196  | 
              raise Fail ("malformed fact identifier " ^ quote ident)
 | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45510 
diff
changeset
 | 
197  | 
end  | 
| 43173 | 198  | 
| NONE => TrueI |> Isa_Raw |> some  | 
199  | 
end  | 
|
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
200  | 
| metis_axiom_from_atp _ _ _ = raise Fail "not CNF -- expected formula"  | 
| 
43092
 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 
blanchet 
parents: 
43091 
diff
changeset
 | 
201  | 
|
| 
45569
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
202  | 
fun eliminate_lam_wrappers (Const (@{const_name Metis.lambda}, _) $ t) =
 | 
| 
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
203  | 
eliminate_lam_wrappers t  | 
| 
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
204  | 
| eliminate_lam_wrappers (t $ u) =  | 
| 
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
205  | 
eliminate_lam_wrappers t $ eliminate_lam_wrappers u  | 
| 
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
206  | 
| eliminate_lam_wrappers (Abs (s, T, t)) =  | 
| 
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
207  | 
Abs (s, T, eliminate_lam_wrappers t)  | 
| 
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
208  | 
| eliminate_lam_wrappers t = t  | 
| 
 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 
blanchet 
parents: 
45565 
diff
changeset
 | 
209  | 
|
| 
39497
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
210  | 
(* Function to generate metis clauses, including comb and type clauses *)  | 
| 45514 | 211  | 
fun prepare_metis_problem ctxt type_enc lam_trans conj_clauses fact_clauses =  | 
| 43212 | 212  | 
let  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
213  | 
val (conj_clauses, fact_clauses) =  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43572 
diff
changeset
 | 
214  | 
if polymorphism_of_type_enc type_enc = Polymorphic then  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
215  | 
(conj_clauses, fact_clauses)  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
216  | 
else  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
217  | 
conj_clauses @ fact_clauses  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
218  | 
|> map (pair 0)  | 
| 
45043
 
7e1a73fc0c8b
drop partial monomorphic instances in Metis, like in Sledgehammer
 
blanchet 
parents: 
45042 
diff
changeset
 | 
219  | 
|> rpair (ctxt |> Config.put Monomorph.keep_partial_instances false)  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
220  | 
|-> Monomorph.monomorph atp_schematic_consts_of  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
221  | 
|> fst |> chop (length conj_clauses)  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
222  | 
|> pairself (maps (map (zero_var_indexes o snd)))  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
223  | 
val num_conjs = length conj_clauses  | 
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
46409 
diff
changeset
 | 
224  | 
(* Pretend every clause is a "simp" rule, to guide the term ordering. *)  | 
| 43212 | 225  | 
val clauses =  | 
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
46409 
diff
changeset
 | 
226  | 
map2 (fn j => pair (Int.toString j, (Local, Simp)))  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
227  | 
(0 upto num_conjs - 1) conj_clauses @  | 
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
46409 
diff
changeset
 | 
228  | 
map2 (fn j => pair (Int.toString (num_conjs + j), (Local, Simp)))  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
229  | 
(0 upto length fact_clauses - 1) fact_clauses  | 
| 43212 | 230  | 
val (old_skolems, props) =  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
231  | 
fold_rev (fn (name, th) => fn (old_skolems, props) =>  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
232  | 
th |> prop_of |> Logic.strip_imp_concl  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
233  | 
|> conceal_old_skolem_terms (length clauses) old_skolems  | 
| 46365 | 234  | 
||> (lam_trans = liftingN orelse lam_trans = lam_liftingN)  | 
235  | 
? eliminate_lam_wrappers  | 
|
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
236  | 
||> (fn prop => (name, prop) :: props))  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
237  | 
clauses ([], [])  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
238  | 
(*  | 
| 
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
239  | 
val _ =  | 
| 
45042
 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 
blanchet 
parents: 
44782 
diff
changeset
 | 
240  | 
      tracing ("PROPS:\n" ^
 | 
| 
 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 
blanchet 
parents: 
44782 
diff
changeset
 | 
241  | 
cat_lines (map (Syntax.string_of_term ctxt o snd) props))  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
242  | 
*)  | 
| 46365 | 243  | 
val lam_trans = if lam_trans = combsN then no_lamsN else lam_trans  | 
| 45551 | 244  | 
val (atp_problem, _, _, lifted, sym_tab) =  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47912 
diff
changeset
 | 
245  | 
prepare_atp_problem ctxt CNF Hypothesis type_enc Metis lam_trans false  | 
| 
47912
 
12de57c5eee5
get rid of "conj_sym_kind" -- most interesting provers now have built-in sorts, and for the others (e.g. E) "Hypothesis" isn't too bad a default
 
blanchet 
parents: 
47046 
diff
changeset
 | 
246  | 
                          false false [] @{prop False} props
 | 
| 45510 | 247  | 
(*  | 
| 
43295
 
30aaab778416
pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
 
blanchet 
parents: 
43268 
diff
changeset
 | 
248  | 
    val _ = tracing ("ATP PROBLEM: " ^
 | 
| 45508 | 249  | 
cat_lines (lines_for_atp_problem CNF atp_problem))  | 
| 45510 | 250  | 
*)  | 
| 45508 | 251  | 
(* "rev" is for compatibility with existing proof scripts. *)  | 
| 43212 | 252  | 
val axioms =  | 
| 
44492
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
253  | 
atp_problem  | 
| 
 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 
blanchet 
parents: 
44411 
diff
changeset
 | 
254  | 
|> maps (map_filter (metis_axiom_from_atp type_enc clauses) o snd) |> rev  | 
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
46409 
diff
changeset
 | 
255  | 
fun ord_info () = atp_problem_term_order_info atp_problem  | 
| 
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
46409 
diff
changeset
 | 
256  | 
in (sym_tab, axioms, ord_info, (lifted, old_skolems)) end  | 
| 
39497
 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 
blanchet 
parents: 
39494 
diff
changeset
 | 
257  | 
|
| 15347 | 258  | 
end;  |