| author | wenzelm | 
| Fri, 06 Sep 2019 15:50:57 +0200 | |
| changeset 70660 | 373d95cf1b98 | 
| parent 69593 | 3dda49e08b9d | 
| child 73932 | fd21b4a93043 | 
| 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: 
39494diff
changeset | 3 | Author: Kong W. Susanto, Cambridge University Computer Laboratory | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
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: 
36378diff
changeset | 5 | Author: Jasmin Blanchette, TU Muenchen | 
| 15347 | 6 | |
| 39494 
bf7dd4902321
rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
 blanchet parents: 
39355diff
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: 
44409diff
changeset | 13 | |
| 43159 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 14 | datatype isa_thm = | 
| 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 15 | Isa_Reflexive_or_Trivial | | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45510diff
changeset | 16 | Isa_Lambda_Lifted | | 
| 43159 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 17 | Isa_Raw of thm | 
| 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
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: 
44411diff
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: 
44411diff
changeset | 23 | val metis_ad_hoc_type_tag : string | 
| 42098 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
41491diff
changeset | 24 | val metis_generated_var_prefix : string | 
| 43231 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
changeset | 25 | val trace : bool Config.T | 
| 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
changeset | 26 | val verbose : bool Config.T | 
| 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
changeset | 27 | val trace_msg : Proof.context -> (unit -> string) -> unit | 
| 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
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: 
44411diff
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: 
39720diff
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: 
45565diff
changeset | 31 | val reveal_lam_lifted : (string * term) list -> term -> term | 
| 57263 | 32 | val generate_metis_problem : Proof.context -> type_enc -> string -> thm list -> thm list -> | 
| 33 | int Symtab.table * (Metis_Thm.thm * isa_thm) list * (unit -> (string * int) list) | |
| 34 | * ((string * term) list * (string * term) list) | |
| 24310 | 35 | end | 
| 15347 | 36 | |
| 46320 | 37 | structure Metis_Generate : METIS_GENERATE = | 
| 15347 | 38 | struct | 
| 39 | ||
| 43092 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43091diff
changeset | 40 | open ATP_Problem | 
| 46320 | 41 | open ATP_Problem_Generate | 
| 15347 | 42 | |
| 43094 | 43 | val metis_equal = "=" | 
| 44 | 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: 
44411diff
changeset | 45 | 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: 
44411diff
changeset | 46 | 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: 
44411diff
changeset | 47 | 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: 
44411diff
changeset | 48 | val metis_ad_hoc_type_tag = "**" | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43003diff
changeset | 49 | val metis_generated_var_prefix = "_" | 
| 17150 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 quigley parents: 
16976diff
changeset | 50 | |
| 69593 | 51 | val trace = Attrib.setup_config_bool \<^binding>\<open>metis_trace\<close> (K false) | 
| 52 | val verbose = Attrib.setup_config_bool \<^binding>\<open>metis_verbose\<close> (K true) | |
| 43231 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
changeset | 53 | |
| 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
changeset | 54 | 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: 
43224diff
changeset | 55 | fun verbose_warning ctxt msg = | 
| 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
changeset | 56 |   if Config.get ctxt verbose then warning ("Metis: " ^ msg) else ()
 | 
| 
69f22ac07395
nicely thread monomorphism verbosity in Metis and Sledgehammer
 blanchet parents: 
43224diff
changeset | 57 | |
| 43094 | 58 | 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: 
44411diff
changeset | 59 | [((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: 
44411diff
changeset | 60 | ((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: 
44411diff
changeset | 61 | ((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: 
44411diff
changeset | 62 | ((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: 
44411diff
changeset | 63 | ((prefixed_type_tag_name, 2), | 
| 44782 | 64 | (fn type_enc => | 
| 65 | if level_of_type_enc type_enc = All_Types then metis_systematic_type_tag | |
| 66 | else metis_ad_hoc_type_tag, true))] | |
| 43094 | 67 | |
| 39896 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 blanchet parents: 
39890diff
changeset | 68 | fun old_skolem_const_name i j num_T_args = | 
| 59877 | 69 | Long_Name.implode (old_skolem_const_prefix :: map string_of_int [i, j, num_T_args]) | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 70 | |
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 71 | fun conceal_old_skolem_terms i old_skolems t = | 
| 69593 | 72 | if exists_Const (curry (op =) \<^const_name>\<open>Meson.skolem\<close> o fst) t then | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 73 | let | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 74 | fun aux old_skolems | 
| 69593 | 75 | (t as (Const (\<^const_name>\<open>Meson.skolem\<close>, Type (_, [_, T])) $ _)) = | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 76 | let | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 77 | val (old_skolems, s) = | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 78 | if i = ~1 then | 
| 69593 | 79 | (old_skolems, \<^const_name>\<open>undefined\<close>) | 
| 57408 | 80 | else | 
| 81 | (case AList.find (op aconv) old_skolems t of | |
| 82 | s :: _ => (old_skolems, s) | |
| 83 | | [] => | |
| 84 | let | |
| 85 | val s = | |
| 86 | old_skolem_const_name i (length old_skolems) (length (Term.add_tvarsT T [])) | |
| 87 | in ((s, t) :: old_skolems, s) end) | |
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 88 | in (old_skolems, Const (s, T)) end | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 89 | | aux old_skolems (t1 $ t2) = | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 90 | let | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 91 | val (old_skolems, t1) = aux old_skolems t1 | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 92 | val (old_skolems, t2) = aux old_skolems t2 | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 93 | in (old_skolems, t1 $ t2) end | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 94 | | aux old_skolems (Abs (s, T, t')) = | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 95 | let val (old_skolems, t') = aux old_skolems t' in | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 96 | (old_skolems, Abs (s, T, t')) | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 97 | end | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 98 | | aux old_skolems t = (old_skolems, t) | 
| 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 99 | in aux old_skolems t end | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 100 | else | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 101 | (old_skolems, t) | 
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 102 | |
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39720diff
changeset | 103 | fun reveal_old_skolem_terms old_skolems = | 
| 37632 | 104 | map_aterms (fn t as Const (s, _) => | 
| 57408 | 105 | if String.isPrefix old_skolem_const_prefix s then | 
| 106 | AList.lookup (op =) old_skolems s |> the | |
| 107 | |> map_types (map_type_tvar (K dummyT)) | |
| 108 | else | |
| 109 | t | |
| 110 | | t => t) | |
| 37632 | 111 | |
| 52150 
41c885784e04
handle lambda-lifted problems in Isar construction code
 blanchet parents: 
52031diff
changeset | 112 | fun reveal_lam_lifted lifted = | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45510diff
changeset | 113 | map_aterms (fn t as Const (s, _) => | 
| 57408 | 114 | if String.isPrefix lam_lifted_prefix s then | 
| 115 | (case AList.lookup (op =) lifted s of | |
| 116 | SOME t => | |
| 69593 | 117 | Const (\<^const_name>\<open>Metis.lambda\<close>, dummyT) | 
| 57408 | 118 | $ map_types (map_type_tvar (K dummyT)) (reveal_lam_lifted lifted t) | 
| 119 | | NONE => t) | |
| 120 | else | |
| 121 | t | |
| 122 | | t => t) | |
| 45508 | 123 | |
| 37577 
5379f41a1322
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
 blanchet parents: 
37575diff
changeset | 124 | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
changeset | 125 | (* ------------------------------------------------------------------------- *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
changeset | 126 | (* Logic maps manage the interface between HOL and first-order logic. *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
changeset | 127 | (* ------------------------------------------------------------------------- *) | 
| 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
changeset | 128 | |
| 43159 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 129 | datatype isa_thm = | 
| 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 130 | Isa_Reflexive_or_Trivial | | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45510diff
changeset | 131 | Isa_Lambda_Lifted | | 
| 43159 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 132 | Isa_Raw of thm | 
| 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 133 | |
| 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 134 | val proxy_defs = map (fst o snd o snd) proxy_table | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53479diff
changeset | 135 | fun prepare_helper ctxt = | 
| 60358 | 136 | Meson.make_meta_clause ctxt #> rewrite_rule ctxt (map safe_mk_meta_eq proxy_defs) | 
| 43159 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43132diff
changeset | 137 | |
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 138 | fun metis_term_of_atp type_enc (ATerm ((s, []), tms)) = | 
| 43094 | 139 | if is_tptp_variable s then | 
| 43268 
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
 blanchet parents: 
43259diff
changeset | 140 | Metis_Term.Var (Metis_Name.fromString s) | 
| 43094 | 141 | else | 
| 44492 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 blanchet parents: 
44411diff
changeset | 142 | (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: 
44411diff
changeset | 143 | 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: 
44411diff
changeset | 144 | | NONE => (s, false)) | 
| 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 blanchet parents: 
44411diff
changeset | 145 | |> (fn (s, swap) => | 
| 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 blanchet parents: 
44411diff
changeset | 146 | Metis_Term.Fn (Metis_Name.fromString s, | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 147 | tms |> map (metis_term_of_atp type_enc) | 
| 44492 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 blanchet parents: 
44411diff
changeset | 148 | |> swap ? rev)) | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 149 | fun metis_atom_of_atp type_enc (AAtom tm) = | 
| 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 150 | (case metis_term_of_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: 
43103diff
changeset | 151 | 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: 
43103diff
changeset | 152 | | _ => raise Fail "non CNF -- expected function") | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 153 | | metis_atom_of_atp _ _ = raise Fail "not CNF -- expected atom" | 
| 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 154 | fun metis_literal_of_atp type_enc (AConn (ANot, [phi])) = | 
| 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 155 | (false, metis_atom_of_atp type_enc phi) | 
| 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 156 | | metis_literal_of_atp type_enc phi = (true, metis_atom_of_atp type_enc phi) | 
| 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 157 | fun metis_literals_of_atp type_enc (AConn (AOr, phis)) = | 
| 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 158 | maps (metis_literals_of_atp type_enc) phis | 
| 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 159 | | metis_literals_of_atp type_enc phi = [metis_literal_of_atp type_enc phi] | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53479diff
changeset | 160 | fun metis_axiom_of_atp ctxt type_enc clauses (Formula ((ident, _), _, phi, _, _)) = | 
| 43173 | 161 | let | 
| 162 | fun some isa = | |
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 163 | SOME (phi |> metis_literals_of_atp type_enc | 
| 44492 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 blanchet parents: 
44411diff
changeset | 164 | |> Metis_LiteralSet.fromList | 
| 43173 | 165 | |> Metis_Thm.axiom, isa) | 
| 166 | in | |
| 47046 
62ca06cc5a99
remove two options that were found to play hardly any role
 blanchet parents: 
47039diff
changeset | 167 | if String.isPrefix tags_sym_formula_prefix ident then | 
| 43173 | 168 | 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: 
43268diff
changeset | 169 | 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: 
43268diff
changeset | 170 | NONE | 
| 43173 | 171 | else if String.isPrefix helper_prefix ident then | 
| 57408 | 172 | (case (String.isSuffix typed_helper_suffix ident, space_explode "_" ident) of | 
| 43194 | 173 | (needs_fairly_sound, _ :: const :: j :: _) => | 
| 57263 | 174 | nth (AList.lookup (op =) helper_table (const, needs_fairly_sound) |> the) | 
| 175 | (the (Int.fromString j) - 1) | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53479diff
changeset | 176 | |> snd |> prepare_helper ctxt | 
| 43194 | 177 | |> Isa_Raw |> some | 
| 57408 | 178 |         | _ => raise Fail ("malformed helper identifier " ^ quote ident))
 | 
| 179 | else | |
| 180 | (case try (unprefix fact_prefix) ident of | |
| 181 | SOME s => | |
| 182 | let val s = s |> space_explode "_" |> tl |> space_implode "_" in | |
| 183 | (case Int.fromString s of | |
| 60358 | 184 | SOME j => Meson.make_meta_clause ctxt (snd (nth clauses j)) |> Isa_Raw |> some | 
| 57408 | 185 | | NONE => | 
| 186 | if String.isPrefix lam_fact_prefix (unascii_of s) then Isa_Lambda_Lifted |> some | |
| 187 |               else raise Fail ("malformed fact identifier " ^ quote ident))
 | |
| 188 | end | |
| 189 | | NONE => some (Isa_Raw TrueI)) | |
| 43173 | 190 | end | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53479diff
changeset | 191 | | metis_axiom_of_atp _ _ _ _ = raise Fail "not CNF -- expected formula" | 
| 43092 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43091diff
changeset | 192 | |
| 69593 | 193 | fun eliminate_lam_wrappers (Const (\<^const_name>\<open>Metis.lambda\<close>, _) $ t) = eliminate_lam_wrappers t | 
| 57263 | 194 | | eliminate_lam_wrappers (t $ u) = eliminate_lam_wrappers t $ eliminate_lam_wrappers u | 
| 195 | | eliminate_lam_wrappers (Abs (s, T, t)) = Abs (s, T, eliminate_lam_wrappers t) | |
| 45569 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 blanchet parents: 
45565diff
changeset | 196 | | eliminate_lam_wrappers t = t | 
| 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 blanchet parents: 
45565diff
changeset | 197 | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
changeset | 198 | (* Function to generate metis clauses, including comb and type clauses *) | 
| 57263 | 199 | fun generate_metis_problem ctxt type_enc lam_trans conj_clauses fact_clauses = | 
| 43212 | 200 | 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: 
43268diff
changeset | 201 | val (conj_clauses, fact_clauses) = | 
| 48131 | 202 | if is_type_enc_polymorphic type_enc 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: 
43268diff
changeset | 203 | (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: 
43268diff
changeset | 204 | 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: 
43268diff
changeset | 205 | 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: 
43268diff
changeset | 206 | |> map (pair 0) | 
| 53479 | 207 | |> Monomorph.monomorph atp_schematic_consts_of ctxt | 
| 208 | |> chop (length conj_clauses) | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
57408diff
changeset | 209 | |> apply2 (maps (map (zero_var_indexes o snd))) | 
| 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: 
43268diff
changeset | 210 | 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: 
46409diff
changeset | 211 | (* Pretend every clause is a "simp" rule, to guide the term ordering. *) | 
| 43212 | 212 | val clauses = | 
| 57263 | 213 | map2 (fn j => pair (Int.toString j, (Local, Simp))) (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: 
46409diff
changeset | 214 | map2 (fn j => pair (Int.toString (num_conjs + j), (Local, Simp))) | 
| 57263 | 215 | (0 upto length fact_clauses - 1) fact_clauses | 
| 43212 | 216 | 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: 
43268diff
changeset | 217 | fold_rev (fn (name, th) => fn (old_skolems, props) => | 
| 59582 | 218 | th |> Thm.prop_of |> Logic.strip_imp_concl | 
| 57263 | 219 | |> conceal_old_skolem_terms (length clauses) old_skolems | 
| 220 | ||> (lam_trans = liftingN orelse lam_trans = lam_liftingN) ? eliminate_lam_wrappers | |
| 221 | ||> (fn prop => (name, prop) :: props)) | |
| 222 | clauses ([], []) | |
| 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: 
43268diff
changeset | 223 | (* | 
| 
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: 
43268diff
changeset | 224 | val _ = | 
| 45042 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44782diff
changeset | 225 |       tracing ("PROPS:\n" ^
 | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44782diff
changeset | 226 | 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: 
43268diff
changeset | 227 | *) | 
| 46365 | 228 | val lam_trans = if lam_trans = combsN then no_lamsN else lam_trans | 
| 57263 | 229 | val (atp_problem, _, lifted, sym_tab) = | 
| 61860 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 blanchet parents: 
60358diff
changeset | 230 | generate_atp_problem ctxt true CNF Hypothesis type_enc Metis lam_trans false false false [] | 
| 69593 | 231 | \<^prop>\<open>False\<close> props | 
| 45510 | 232 | (* | 
| 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: 
43268diff
changeset | 233 |     val _ = tracing ("ATP PROBLEM: " ^
 | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51575diff
changeset | 234 | cat_lines (lines_of_atp_problem CNF atp_problem)) | 
| 45510 | 235 | *) | 
| 45508 | 236 | (* "rev" is for compatibility with existing proof scripts. *) | 
| 57263 | 237 | val axioms = atp_problem | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53479diff
changeset | 238 | |> maps (map_filter (metis_axiom_of_atp ctxt 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: 
46409diff
changeset | 239 | fun ord_info () = atp_problem_term_order_info atp_problem | 
| 57408 | 240 | in | 
| 241 | (sym_tab, axioms, ord_info, (lifted, old_skolems)) | |
| 242 | end | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
changeset | 243 | |
| 15347 | 244 | end; |