author | blanchet |
Sun, 13 Nov 2011 20:28:22 +0100 | |
changeset 45478 | 8e299034eab4 |
parent 44492 | a330c0608da8 |
child 45508 | b216dc1b3630 |
permissions | -rw-r--r-- |
39958 | 1 |
(* Title: HOL/Tools/Metis/metis_reconstruct.ML |
39495
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
2 |
Author: Kong W. Susanto, Cambridge University Computer Laboratory |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
3 |
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
5 |
Copyright Cambridge University 2007 |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
6 |
|
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
7 |
Proof reconstruction for Metis. |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
8 |
*) |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
9 |
|
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
10 |
signature METIS_RECONSTRUCT = |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
11 |
sig |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
12 |
type type_enc = ATP_Translate.type_enc |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
13 |
|
42650
552eae49f97d
reintroduce this idea of running "metisFT" after a failed "metis" -- I took it out in e85ce10cef1a because I couldn't think of a reasonable use case, but now that ATPs use sound encodings and include dangerous facts (e.g. True_or_False) it makes more sense than ever to run "metisFT" after "metis"
blanchet
parents:
42616
diff
changeset
|
14 |
exception METIS of string * string |
552eae49f97d
reintroduce this idea of running "metisFT" after a failed "metis" -- I took it out in e85ce10cef1a because I couldn't think of a reasonable use case, but now that ATPs use sound encodings and include dangerous facts (e.g. True_or_False) it makes more sense than ever to run "metisFT" after "metis"
blanchet
parents:
42616
diff
changeset
|
15 |
|
43159
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents:
43136
diff
changeset
|
16 |
val hol_clause_from_metis : |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
17 |
Proof.context -> type_enc -> int Symtab.table -> (string * term) list |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
18 |
-> Metis_Thm.thm -> term |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
19 |
val lookth : (Metis_Thm.thm * 'a) list -> Metis_Thm.thm -> 'a |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
20 |
val replay_one_inference : |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
21 |
Proof.context -> type_enc -> (string * term) list -> int Symtab.table |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
22 |
-> Metis_Thm.thm * Metis_Proof.inference -> (Metis_Thm.thm * thm) list |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
23 |
-> (Metis_Thm.thm * thm) list |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
24 |
val discharge_skolem_premises : |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
25 |
Proof.context -> (thm * term) option list -> thm -> thm |
39495
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
26 |
end; |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
27 |
|
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
28 |
structure Metis_Reconstruct : METIS_RECONSTRUCT = |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
29 |
struct |
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
30 |
|
43094 | 31 |
open ATP_Problem |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
42650
diff
changeset
|
32 |
open ATP_Translate |
43094 | 33 |
open ATP_Reconstruct |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
34 |
open Metis_Translate |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
35 |
|
42650
552eae49f97d
reintroduce this idea of running "metisFT" after a failed "metis" -- I took it out in e85ce10cef1a because I couldn't think of a reasonable use case, but now that ATPs use sound encodings and include dangerous facts (e.g. True_or_False) it makes more sense than ever to run "metisFT" after "metis"
blanchet
parents:
42616
diff
changeset
|
36 |
exception METIS of string * string |
552eae49f97d
reintroduce this idea of running "metisFT" after a failed "metis" -- I took it out in e85ce10cef1a because I couldn't think of a reasonable use case, but now that ATPs use sound encodings and include dangerous facts (e.g. True_or_False) it makes more sense than ever to run "metisFT" after "metis"
blanchet
parents:
42616
diff
changeset
|
37 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
38 |
fun atp_name_from_metis type_enc s = |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
39 |
case find_first (fn (_, (f, _)) => f type_enc = s) metis_name_table 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
|
40 |
SOME ((s, _), (_, swap)) => (s, swap) |
81d1b15aa0ae
use ":" for type information (looks good in Metis's output) and handle it in new path finder
blanchet
parents:
43103
diff
changeset
|
41 |
| _ => (s, false) |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
42 |
fun atp_term_from_metis type_enc (Metis_Term.Fn (s, tms)) = |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
43 |
let val (s, swap) = atp_name_from_metis type_enc (Metis_Name.toString s) in |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
44 |
ATerm (s, tms |> map (atp_term_from_metis type_enc) |> swap ? rev) |
43104
81d1b15aa0ae
use ":" for type information (looks good in Metis's output) and handle it in new path finder
blanchet
parents:
43103
diff
changeset
|
45 |
end |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
46 |
| atp_term_from_metis _ (Metis_Term.Var s) = ATerm (Metis_Name.toString s, []) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
47 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
48 |
fun hol_term_from_metis ctxt type_enc sym_tab = |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
49 |
atp_term_from_metis type_enc #> term_from_atp ctxt false sym_tab NONE |
43094 | 50 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
51 |
fun atp_literal_from_metis type_enc (pos, atom) = |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
52 |
atom |> Metis_Term.Fn |> atp_term_from_metis type_enc |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
53 |
|> AAtom |> not pos ? mk_anot |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
54 |
fun atp_clause_from_metis _ [] = AAtom (ATerm (tptp_false, [])) |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
55 |
| atp_clause_from_metis type_enc lits = |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
56 |
lits |> map (atp_literal_from_metis type_enc) |> mk_aconns AOr |
43136
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents:
43135
diff
changeset
|
57 |
|
43184 | 58 |
fun reveal_old_skolems_and_infer_types ctxt old_skolems = |
59 |
map (reveal_old_skolem_terms old_skolems) |
|
43212 | 60 |
#> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_pattern ctxt) |
43184 | 61 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
62 |
fun hol_clause_from_metis ctxt type_enc sym_tab old_skolems = |
43159
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents:
43136
diff
changeset
|
63 |
Metis_Thm.clause |
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents:
43136
diff
changeset
|
64 |
#> Metis_LiteralSet.toList |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
65 |
#> atp_clause_from_metis type_enc |
43159
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents:
43136
diff
changeset
|
66 |
#> prop_from_atp ctxt false sym_tab |
43184 | 67 |
#> singleton (reveal_old_skolems_and_infer_types ctxt old_skolems) |
43136
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents:
43135
diff
changeset
|
68 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
69 |
fun hol_terms_from_metis ctxt type_enc old_skolems sym_tab fol_tms = |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
70 |
let val ts = map (hol_term_from_metis ctxt type_enc sym_tab) fol_tms |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
71 |
val _ = trace_msg ctxt (fn () => " calling type inference:") |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
72 |
val _ = app (fn t => trace_msg ctxt |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
73 |
(fn () => Syntax.string_of_term ctxt t)) ts |
43184 | 74 |
val ts' = |
75 |
ts |> reveal_old_skolems_and_infer_types ctxt old_skolems |
|
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
76 |
val _ = app (fn t => trace_msg ctxt |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
77 |
(fn () => " final term: " ^ Syntax.string_of_term ctxt t ^ |
43128 | 78 |
" of type " ^ Syntax.string_of_typ ctxt (type_of t))) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
79 |
ts' |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
80 |
in ts' end; |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
81 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
82 |
(* ------------------------------------------------------------------------- *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
83 |
(* FOL step Inference Rules *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
84 |
(* ------------------------------------------------------------------------- *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
85 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
86 |
(*for debugging only*) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
87 |
(* |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
88 |
fun print_thpair ctxt (fth,th) = |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
89 |
(trace_msg ctxt (fn () => "============================================="); |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
90 |
trace_msg ctxt (fn () => "Metis: " ^ Metis_Thm.toString fth); |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
91 |
trace_msg ctxt (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th)); |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
92 |
*) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
93 |
|
43094 | 94 |
fun lookth th_pairs fth = |
95 |
the (AList.lookup (uncurry Metis_Thm.equal) th_pairs fth) |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
96 |
handle Option.Option => |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
97 |
raise Fail ("Failed to find Metis theorem " ^ Metis_Thm.toString fth) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
98 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
99 |
fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx)); |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
100 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
101 |
(* INFERENCE RULE: AXIOM *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
102 |
|
43212 | 103 |
(* This causes variables to have an index of 1 by default. See also |
104 |
"term_from_atp" in "ATP_Reconstruct". *) |
|
105 |
val axiom_inference = Thm.incr_indexes 1 oo lookth |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
106 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
107 |
(* INFERENCE RULE: ASSUME *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
108 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
109 |
val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)} |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
110 |
|
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
111 |
fun inst_excluded_middle thy i_atom = |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
112 |
let val th = EXCLUDED_MIDDLE |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
113 |
val [vx] = Term.add_vars (prop_of th) [] |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
114 |
val substs = [(cterm_of thy (Var vx), cterm_of thy i_atom)] |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
115 |
in cterm_instantiate substs th end; |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
116 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
117 |
fun assume_inference ctxt type_enc old_skolems sym_tab atom = |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
118 |
inst_excluded_middle |
42361 | 119 |
(Proof_Context.theory_of ctxt) |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
120 |
(singleton (hol_terms_from_metis ctxt type_enc old_skolems sym_tab) |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
121 |
(Metis_Term.Fn atom)) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
122 |
|
39498
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
123 |
(* INFERENCE RULE: INSTANTIATE (SUBST). Type instantiations are ignored. Trying |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
124 |
to reconstruct them admits new possibilities of errors, e.g. concerning |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
125 |
sorts. Instead we try to arrange that new TVars are distinct and that types |
39498
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
126 |
can be inferred from terms. *) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
127 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
128 |
fun inst_inference ctxt type_enc old_skolems sym_tab th_pairs fsubst th = |
42361 | 129 |
let val thy = Proof_Context.theory_of ctxt |
43094 | 130 |
val i_th = lookth th_pairs th |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
131 |
val i_th_vars = Term.add_vars (prop_of i_th) [] |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
132 |
fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
133 |
fun subst_translation (x,y) = |
39498
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
134 |
let val v = find_var x |
43184 | 135 |
(* We call "reveal_old_skolems_and_infer_types" below. *) |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
136 |
val t = hol_term_from_metis ctxt type_enc sym_tab y |
39498
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
137 |
in SOME (cterm_of thy (Var v), t) end |
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
138 |
handle Option.Option => |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
139 |
(trace_msg ctxt (fn () => "\"find_var\" failed for " ^ x ^ |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
140 |
" in " ^ Display.string_of_thm ctxt i_th); |
39498
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
141 |
NONE) |
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
142 |
| TYPE _ => |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
143 |
(trace_msg ctxt (fn () => "\"hol_term_from_metis\" failed for " ^ x ^ |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
144 |
" in " ^ Display.string_of_thm ctxt i_th); |
39498
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
145 |
NONE) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
146 |
fun remove_typeinst (a, t) = |
43268
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
147 |
let val a = Metis_Name.toString a in |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
148 |
case strip_prefix_and_unascii schematic_var_prefix a of |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
149 |
SOME b => SOME (b, t) |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
150 |
| NONE => |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
151 |
case strip_prefix_and_unascii tvar_prefix a of |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
152 |
SOME _ => NONE (* type instantiations are forbidden *) |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
153 |
| NONE => SOME (a, t) (* internal Metis var? *) |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
154 |
end |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
155 |
val _ = trace_msg ctxt (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
156 |
val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst) |
43184 | 157 |
val (vars, tms) = |
158 |
ListPair.unzip (map_filter subst_translation substs) |
|
159 |
||> reveal_old_skolems_and_infer_types ctxt old_skolems |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
160 |
val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
161 |
val substs' = ListPair.zip (vars, map ctm_of tms) |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
162 |
val _ = trace_msg ctxt (fn () => |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
163 |
cat_lines ("subst_translations:" :: |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
164 |
(substs' |> map (fn (x, y) => |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
165 |
Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^ |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
166 |
Syntax.string_of_term ctxt (term_of y))))); |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
167 |
in cterm_instantiate substs' i_th end |
43186 | 168 |
handle THM (msg, _, _) => raise METIS ("inst_inference", msg) |
169 |
| ERROR msg => raise METIS ("inst_inference", msg) |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
170 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
171 |
(* INFERENCE RULE: RESOLVE *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
172 |
|
43330 | 173 |
(*Increment the indexes of only the type variables*) |
174 |
fun incr_type_indexes inc th = |
|
175 |
let val tvs = Term.add_tvars (Thm.full_prop_of th) [] |
|
176 |
and thy = Thm.theory_of_thm th |
|
177 |
fun inc_tvar ((a,i),s) = pairself (ctyp_of thy) (TVar ((a,i),s), TVar ((a,i+inc),s)) |
|
178 |
in Thm.instantiate (map inc_tvar tvs, []) th end; |
|
179 |
||
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
180 |
(* Like RSN, but we rename apart only the type variables. Vars here typically |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
181 |
have an index of 1, and the use of RSN would increase this typically to 3. |
43300 | 182 |
Instantiations of those Vars could then fail. *) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
183 |
fun resolve_inc_tyvars thy tha i thb = |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
184 |
let |
43330 | 185 |
val tha = incr_type_indexes (1 + Thm.maxidx_of thb) tha |
43359 | 186 |
fun aux (tha, thb) = |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
187 |
case Thm.bicompose false (false, tha, nprems_of tha) i thb |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
188 |
|> Seq.list_of |> distinct Thm.eq_thm of |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
189 |
[th] => th |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
190 |
| _ => raise THM ("resolve_inc_tyvars: unique result expected", i, |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
191 |
[tha, thb]) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
192 |
in |
43359 | 193 |
aux (tha, thb) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
194 |
handle TERM z => |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
195 |
(* The unifier, which is invoked from "Thm.bicompose", will sometimes |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
196 |
refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
197 |
"TERM" exception (with "add_ffpair" as first argument). We then |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
198 |
perform unification of the types of variables by hand and try |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
199 |
again. We could do this the first time around but this error |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
200 |
occurs seldom and we don't want to break existing proofs in subtle |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
201 |
ways or slow them down needlessly. *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
202 |
case [] |> fold (Term.add_vars o prop_of) [tha, thb] |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
203 |
|> AList.group (op =) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
204 |
|> maps (fn ((s, _), T :: Ts) => |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
205 |
map (fn T' => (Free (s, T), Free (s, T'))) Ts) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
206 |
|> rpair (Envir.empty ~1) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
207 |
|-> fold (Pattern.unify thy) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
208 |
|> Envir.type_env |> Vartab.dest |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
209 |
|> map (fn (x, (S, T)) => |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
210 |
pairself (ctyp_of thy) (TVar (x, S), T)) of |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
211 |
[] => raise TERM z |
43359 | 212 |
| ps => (tha, thb) |> pairself (Drule.instantiate_normalize (ps, [])) |
213 |
|> aux |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
214 |
end |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
215 |
|
40221
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
blanchet
parents:
40158
diff
changeset
|
216 |
fun s_not (@{const Not} $ t) = t |
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
blanchet
parents:
40158
diff
changeset
|
217 |
| s_not t = HOLogic.mk_not t |
43195
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
218 |
fun simp_not_not (@{const Trueprop} $ t) = @{const Trueprop} $ simp_not_not t |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
219 |
| simp_not_not (@{const Not} $ t) = s_not (simp_not_not t) |
40221
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
blanchet
parents:
40158
diff
changeset
|
220 |
| simp_not_not t = t |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
221 |
|
43195
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
222 |
val normalize_literal = simp_not_not o Envir.eta_contract |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
223 |
|
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
224 |
(* Find the relative location of an untyped term within a list of terms as a |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
225 |
1-based index. Returns 0 in case of failure. *) |
40221
d10b68c6e6d4
do not let Metis be confused by higher-order reasoning leading to literals of the form "~ ~ p", which are really the same as "p"
blanchet
parents:
40158
diff
changeset
|
226 |
fun index_of_literal lit haystack = |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
227 |
let |
43195
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
228 |
fun match_lit normalize = |
43134
0c82e00ba63e
make sure no warnings are given for polymorphic facts where we use a monomorphic instance
blanchet
parents:
43130
diff
changeset
|
229 |
HOLogic.dest_Trueprop #> normalize |
43301
8d7fc4a5b502
removed needless function that duplicated standard functionality, with a little unnecessary twist
blanchet
parents:
43300
diff
changeset
|
230 |
#> curry Term.aconv_untyped (lit |> normalize) |
43195
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
231 |
in |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
232 |
(case find_index (match_lit I) haystack of |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
233 |
~1 => find_index (match_lit (simp_not_not o Envir.eta_contract)) haystack |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
234 |
| j => j) + 1 |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
235 |
end |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
236 |
|
39893 | 237 |
(* Permute a rule's premises to move the i-th premise to the last position. *) |
238 |
fun make_last i th = |
|
239 |
let val n = nprems_of th |
|
240 |
in if 1 <= i andalso i <= n |
|
241 |
then Thm.permute_prems (i-1) 1 th |
|
242 |
else raise THM("select_literal", i, [th]) |
|
243 |
end; |
|
244 |
||
42348
187354e22c7d
improve on 0b05cc14c2cb: make sure that a literal variable "?foo" isn't accidentally renamed "?Q", which might be enough to confuse the new Skolemizer (cf. "Clausify.thy" example)
blanchet
parents:
42344
diff
changeset
|
245 |
(* Maps a rule that ends "... ==> P ==> False" to "... ==> ~ P" while avoiding |
42349
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
246 |
to create double negations. The "select" wrapper is a trick to ensure that |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
247 |
"P ==> ~ False ==> False" is rewritten to "P ==> False", not to "~ P". We |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
248 |
don't use this trick in general because it makes the proof object uglier than |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
249 |
necessary. FIXME. *) |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
250 |
fun negate_head th = |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
251 |
if exists (fn t => t aconv @{prop "~ False"}) (prems_of th) then |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
252 |
(th RS @{thm select_FalseI}) |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
253 |
|> fold (rewrite_rule o single) |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
254 |
@{thms not_atomize_select atomize_not_select} |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
255 |
else |
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
blanchet
parents:
42348
diff
changeset
|
256 |
th |> fold (rewrite_rule o single) @{thms not_atomize atomize_not} |
39893 | 257 |
|
258 |
(* Maps the clause [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P *) |
|
42348
187354e22c7d
improve on 0b05cc14c2cb: make sure that a literal variable "?foo" isn't accidentally renamed "?Q", which might be enough to confuse the new Skolemizer (cf. "Clausify.thy" example)
blanchet
parents:
42344
diff
changeset
|
259 |
val select_literal = negate_head oo make_last |
39893 | 260 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
261 |
fun resolve_inference ctxt type_enc old_skolems sym_tab th_pairs atom th1 th2 = |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
262 |
let |
43094 | 263 |
val (i_th1, i_th2) = pairself (lookth th_pairs) (th1, th2) |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
264 |
val _ = trace_msg ctxt (fn () => |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
265 |
" isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1 ^ "\n\ |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
266 |
\ isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
267 |
in |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
268 |
(* Trivial cases where one operand is type info *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
269 |
if Thm.eq_thm (TrueI, i_th1) then |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
270 |
i_th2 |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
271 |
else if Thm.eq_thm (TrueI, i_th2) then |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
272 |
i_th1 |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
273 |
else |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
274 |
let |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
275 |
val thy = Proof_Context.theory_of ctxt |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
276 |
val i_atom = |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
277 |
singleton (hol_terms_from_metis ctxt type_enc old_skolems sym_tab) |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
278 |
(Metis_Term.Fn atom) |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
279 |
val _ = trace_msg ctxt (fn () => |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
280 |
" atom: " ^ Syntax.string_of_term ctxt i_atom) |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
281 |
in |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
282 |
case index_of_literal (s_not i_atom) (prems_of i_th1) of |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
283 |
0 => |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
284 |
(trace_msg ctxt (fn () => "Failed to find literal in \"th1\""); |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
285 |
i_th1) |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
286 |
| j1 => |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
287 |
(trace_msg ctxt (fn () => " index th1: " ^ string_of_int j1); |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
288 |
case index_of_literal i_atom (prems_of i_th2) of |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
289 |
0 => |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
290 |
(trace_msg ctxt (fn () => "Failed to find literal in \"th2\""); |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
291 |
i_th2) |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
292 |
| j2 => |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
293 |
(trace_msg ctxt (fn () => " index th2: " ^ string_of_int j2); |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
294 |
resolve_inc_tyvars thy (select_literal j1 i_th1) j2 i_th2 |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
295 |
handle TERM (s, _) => raise METIS ("resolve_inference", s))) |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
296 |
end |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
297 |
end |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
298 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
299 |
(* INFERENCE RULE: REFL *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
300 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
301 |
val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp} |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
302 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
303 |
val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) []))); |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
304 |
val refl_idx = 1 + Thm.maxidx_of REFL_THM; |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
305 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
306 |
fun refl_inference ctxt type_enc old_skolems sym_tab t = |
43094 | 307 |
let |
308 |
val thy = Proof_Context.theory_of ctxt |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
309 |
val i_t = |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
310 |
singleton (hol_terms_from_metis ctxt type_enc old_skolems sym_tab) t |
43094 | 311 |
val _ = trace_msg ctxt (fn () => " term: " ^ Syntax.string_of_term ctxt i_t) |
312 |
val c_t = cterm_incr_types thy refl_idx i_t |
|
313 |
in cterm_instantiate [(refl_x, c_t)] REFL_THM end |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
314 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
315 |
(* INFERENCE RULE: EQUALITY *) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
316 |
|
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
317 |
val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp} |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
318 |
val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp} |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
319 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
320 |
fun equality_inference ctxt type_enc old_skolems sym_tab (pos, atom) fp fr = |
42361 | 321 |
let val thy = Proof_Context.theory_of ctxt |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
322 |
val m_tm = Metis_Term.Fn atom |
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
323 |
val [i_atom, i_tm] = |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
324 |
hol_terms_from_metis ctxt type_enc old_skolems sym_tab [m_tm, fr] |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
325 |
val _ = trace_msg ctxt (fn () => "sign of the literal: " ^ Bool.toString pos) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
326 |
fun replace_item_list lx 0 (_::ls) = lx::ls |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
327 |
| replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls |
43205 | 328 |
fun path_finder_fail tm ps t = |
43209
007117fed183
fall back in case path finder fails -- these errors are sometimes salvageable
blanchet
parents:
43205
diff
changeset
|
329 |
raise METIS ("equality_inference (path_finder)", |
007117fed183
fall back in case path finder fails -- these errors are sometimes salvageable
blanchet
parents:
43205
diff
changeset
|
330 |
"path = " ^ space_implode " " (map string_of_int ps) ^ |
007117fed183
fall back in case path finder fails -- these errors are sometimes salvageable
blanchet
parents:
43205
diff
changeset
|
331 |
" isa-term: " ^ Syntax.string_of_term ctxt tm ^ |
007117fed183
fall back in case path finder fails -- these errors are sometimes salvageable
blanchet
parents:
43205
diff
changeset
|
332 |
(case t of |
007117fed183
fall back in case path finder fails -- these errors are sometimes salvageable
blanchet
parents:
43205
diff
changeset
|
333 |
SOME t => " fol-term: " ^ Metis_Term.toString t |
007117fed183
fall back in case path finder fails -- these errors are sometimes salvageable
blanchet
parents:
43205
diff
changeset
|
334 |
| NONE => "")) |
43212 | 335 |
fun path_finder tm [] _ = (tm, Bound 0) |
336 |
| path_finder tm (p :: ps) (t as Metis_Term.Fn (s, ts)) = |
|
43177 | 337 |
let |
43268
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
338 |
val s = s |> Metis_Name.toString |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
339 |
|> perhaps (try (strip_prefix_and_unascii const_prefix |
43177 | 340 |
#> the #> unmangled_const_name)) |
341 |
in |
|
342 |
if s = metis_predicator orelse s = predicator_name orelse |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
343 |
s = metis_systematic_type_tag orelse s = metis_ad_hoc_type_tag |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
344 |
orelse s = type_tag_name then |
43212 | 345 |
path_finder tm ps (nth ts p) |
43177 | 346 |
else if s = metis_app_op orelse s = app_op_name then |
43130
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
347 |
let |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
348 |
val (tm1, tm2) = dest_comb tm |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
349 |
val p' = p - (length ts - 2) |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
350 |
in |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
351 |
if p' = 0 then |
43212 | 352 |
path_finder tm1 ps (nth ts p) ||> (fn y => y $ tm2) |
43130
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
353 |
else |
43212 | 354 |
path_finder tm2 ps (nth ts p) ||> (fn y => tm1 $ y) |
43130
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
355 |
end |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
356 |
else |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
357 |
let |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
358 |
val (tm1, args) = strip_comb tm |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
359 |
val adjustment = length ts - length args |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
360 |
val p' = if adjustment > p then p else p - adjustment |
43205 | 361 |
val tm_p = |
362 |
nth args p' |
|
43278 | 363 |
handle General.Subscript => path_finder_fail tm (p :: ps) (SOME t) |
43130
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
364 |
val _ = trace_msg ctxt (fn () => |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
365 |
"path_finder: " ^ string_of_int p ^ " " ^ |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
366 |
Syntax.string_of_term ctxt tm_p) |
43212 | 367 |
val (r, t) = path_finder tm_p ps (nth ts p) |
43130
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
368 |
in (r, list_comb (tm1, replace_item_list t p' args)) end |
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
blanchet
parents:
43128
diff
changeset
|
369 |
end |
43212 | 370 |
| path_finder tm ps t = path_finder_fail tm ps (SOME t) |
371 |
val (tm_subst, body) = path_finder i_atom fp m_tm |
|
39498
e8aef7ea9cbb
make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents:
39497
diff
changeset
|
372 |
val tm_abs = Abs ("x", type_of tm_subst, body) |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
373 |
val _ = trace_msg ctxt (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs) |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
374 |
val _ = trace_msg ctxt (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm) |
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
375 |
val _ = trace_msg ctxt (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
376 |
val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill typed but gives right max*) |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
377 |
val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em) |
39978
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
blanchet
parents:
39964
diff
changeset
|
378 |
val _ = trace_msg ctxt (fn () => "subst' " ^ Display.string_of_thm ctxt subst') |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
379 |
val eq_terms = map (pairself (cterm_of thy)) |
44121 | 380 |
(ListPair.zip (Misc_Legacy.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm])) |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
381 |
in cterm_instantiate eq_terms subst' end; |
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
382 |
|
43094 | 383 |
val factor = Seq.hd o distinct_subgoals_tac |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
384 |
|
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
385 |
fun one_step ctxt type_enc old_skolems sym_tab th_pairs p = |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
386 |
case p of |
43186 | 387 |
(fol_th, Metis_Proof.Axiom _) => axiom_inference th_pairs fol_th |> factor |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
388 |
| (_, Metis_Proof.Assume f_atom) => |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
389 |
assume_inference ctxt type_enc old_skolems sym_tab f_atom |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
390 |
| (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) => |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
391 |
inst_inference ctxt type_enc old_skolems sym_tab th_pairs f_subst f_th1 |
43186 | 392 |
|> factor |
43187
95bd1ef1331a
make resolution replay more robust, in case Metis distinguishes between two literals that are merged in Isabelle (e.g. because they carry more or less type annotations in Metis)
blanchet
parents:
43186
diff
changeset
|
393 |
| (_, Metis_Proof.Resolve(f_atom, f_th1, f_th2)) => |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
394 |
resolve_inference ctxt type_enc old_skolems sym_tab th_pairs f_atom f_th1 |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
395 |
f_th2 |
43094 | 396 |
|> factor |
43186 | 397 |
| (_, Metis_Proof.Refl f_tm) => |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
398 |
refl_inference ctxt type_enc old_skolems sym_tab f_tm |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
399 |
| (_, Metis_Proof.Equality (f_lit, f_p, f_r)) => |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
400 |
equality_inference ctxt type_enc old_skolems sym_tab f_lit f_p f_r |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
401 |
|
39893 | 402 |
fun flexflex_first_order th = |
403 |
case Thm.tpairs_of th of |
|
404 |
[] => th |
|
405 |
| pairs => |
|
406 |
let val thy = theory_of_thm th |
|
407 |
val (_, tenv) = |
|
408 |
fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty) |
|
409 |
val t_pairs = map Meson.term_pair_of (Vartab.dest tenv) |
|
410 |
val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th |
|
411 |
in th' end |
|
412 |
handle THM _ => th; |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
413 |
|
43268
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
414 |
fun is_metis_literal_genuine (_, (s, _)) = |
c0eaa8b9bff5
removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents:
43262
diff
changeset
|
415 |
not (String.isPrefix class_prefix (Metis_Name.toString s)) |
39895
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents:
39893
diff
changeset
|
416 |
fun is_isabelle_literal_genuine t = |
39953
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents:
39946
diff
changeset
|
417 |
case t of _ $ (Const (@{const_name Meson.skolem}, _) $ _) => false | _ => true |
39895
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents:
39893
diff
changeset
|
418 |
|
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents:
39893
diff
changeset
|
419 |
fun count p xs = fold (fn x => if p x then Integer.add 1 else I) xs 0 |
a91a84b1dfdd
reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents:
39893
diff
changeset
|
420 |
|
42333 | 421 |
(* Seldomly needed hack. A Metis clause is represented as a set, so duplicate |
422 |
disjuncts are impossible. In the Isabelle proof, in spite of efforts to |
|
423 |
eliminate them, duplicates sometimes appear with slightly different (but |
|
424 |
unifiable) types. *) |
|
425 |
fun resynchronize ctxt fol_th th = |
|
426 |
let |
|
427 |
val num_metis_lits = |
|
428 |
count is_metis_literal_genuine |
|
429 |
(Metis_LiteralSet.toList (Metis_Thm.clause fol_th)) |
|
430 |
val num_isabelle_lits = count is_isabelle_literal_genuine (prems_of th) |
|
431 |
in |
|
432 |
if num_metis_lits >= num_isabelle_lits then |
|
433 |
th |
|
434 |
else |
|
435 |
let |
|
436 |
val (prems0, concl) = th |> prop_of |> Logic.strip_horn |
|
43195
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
437 |
val prems = prems0 |> map normalize_literal |
43301
8d7fc4a5b502
removed needless function that duplicated standard functionality, with a little unnecessary twist
blanchet
parents:
43300
diff
changeset
|
438 |
|> distinct Term.aconv_untyped |
42333 | 439 |
val goal = Logic.list_implies (prems, concl) |
43195
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
440 |
val tac = cut_rules_tac [th] 1 |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
441 |
THEN rewrite_goals_tac @{thms not_not [THEN eq_reflection]} |
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
442 |
THEN ALLGOALS assume_tac |
42333 | 443 |
in |
444 |
if length prems = length prems0 then |
|
42650
552eae49f97d
reintroduce this idea of running "metisFT" after a failed "metis" -- I took it out in e85ce10cef1a because I couldn't think of a reasonable use case, but now that ATPs use sound encodings and include dangerous facts (e.g. True_or_False) it makes more sense than ever to run "metisFT" after "metis"
blanchet
parents:
42616
diff
changeset
|
445 |
raise METIS ("resynchronize", "Out of sync") |
42333 | 446 |
else |
43195
6dc58b3b73b5
improved correctness of handling of higher-order occurrences of "Not" in new Metis (and probably in old Metis)
blanchet
parents:
43187
diff
changeset
|
447 |
Goal.prove ctxt [] [] goal (K tac) |
42333 | 448 |
|> resynchronize ctxt fol_th |
449 |
end |
|
450 |
end |
|
451 |
||
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
452 |
fun replay_one_inference ctxt type_enc old_skolems sym_tab (fol_th, inf) |
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
453 |
th_pairs = |
43094 | 454 |
if not (null th_pairs) andalso |
455 |
prop_of (snd (hd th_pairs)) aconv @{prop False} then |
|
40868
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
456 |
(* Isabelle sometimes identifies literals (premises) that are distinct in |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
457 |
Metis (e.g., because of type variables). We give the Isabelle proof the |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
458 |
benefice of the doubt. *) |
43094 | 459 |
th_pairs |
40868
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
460 |
else |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
461 |
let |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
462 |
val _ = trace_msg ctxt |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
463 |
(fn () => "=============================================") |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
464 |
val _ = trace_msg ctxt |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
465 |
(fn () => "METIS THM: " ^ Metis_Thm.toString fol_th) |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
466 |
val _ = trace_msg ctxt |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
467 |
(fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf) |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44241
diff
changeset
|
468 |
val th = one_step ctxt type_enc old_skolems sym_tab th_pairs (fol_th, inf) |
40868
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
469 |
|> flexflex_first_order |
42333 | 470 |
|> resynchronize ctxt fol_th |
40868
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
471 |
val _ = trace_msg ctxt |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
472 |
(fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th) |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
473 |
val _ = trace_msg ctxt |
177cd660abb7
give the Isabelle proof the benefice of the doubt when the Isabelle theorem has fewer literals than the Metis one -- this makes a difference on lemma "Let (x::'a, y::'a) (inv_image (r::'b * 'b => bool) (f::'a => 'b)) = ((f x, f y) : r)" apply (metis in_inv_image mem_def)
blanchet
parents:
40724
diff
changeset
|
474 |
(fn () => "=============================================") |
43094 | 475 |
in (fol_th, th) :: th_pairs end |
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39495
diff
changeset
|
476 |
|
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
477 |
(* It is normally sufficient to apply "assume_tac" to unify the conclusion with |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
478 |
one of the premises. Unfortunately, this sometimes yields "Variable |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
479 |
?SK_a_b_c_x has two distinct types" errors. To avoid this, we instantiate the |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
480 |
variables before applying "assume_tac". Typical constraints are of the form |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
481 |
?SK_a_b_c_x SK_d_e_f_y ... SK_a_b_c_x ... SK_g_h_i_z =?= SK_a_b_c_x, |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
482 |
where the nonvariables are goal parameters. *) |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
483 |
fun unify_first_prem_with_concl thy i th = |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
484 |
let |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
485 |
val goal = Logic.get_goal (prop_of th) i |> Envir.beta_eta_contract |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
486 |
val prem = goal |> Logic.strip_assums_hyp |> hd |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
487 |
val concl = goal |> Logic.strip_assums_concl |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
488 |
fun pair_untyped_aconv (t1, t2) (u1, u2) = |
43301
8d7fc4a5b502
removed needless function that duplicated standard functionality, with a little unnecessary twist
blanchet
parents:
43300
diff
changeset
|
489 |
Term.aconv_untyped (t1, u1) andalso Term.aconv_untyped (t2, u2) |
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
490 |
fun add_terms tp inst = |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
491 |
if exists (pair_untyped_aconv tp) inst then inst |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
492 |
else tp :: map (apsnd (subst_atomic [tp])) inst |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
493 |
fun is_flex t = |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
494 |
case strip_comb t of |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
495 |
(Var _, args) => forall is_Bound args |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
496 |
| _ => false |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
497 |
fun unify_flex flex rigid = |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
498 |
case strip_comb flex of |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
499 |
(Var (z as (_, T)), args) => |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
500 |
add_terms (Var z, |
44241 | 501 |
fold_rev absdummy (take (length args) (binder_types T)) rigid) |
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
502 |
| _ => I |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
503 |
fun unify_potential_flex comb atom = |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
504 |
if is_flex comb then unify_flex comb atom |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
505 |
else if is_Var atom then add_terms (atom, comb) |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
506 |
else I |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
507 |
fun unify_terms (t, u) = |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
508 |
case (t, u) of |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
509 |
(t1 $ t2, u1 $ u2) => |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
510 |
if is_flex t then unify_flex t u |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
511 |
else if is_flex u then unify_flex u t |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
512 |
else fold unify_terms [(t1, u1), (t2, u2)] |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
513 |
| (_ $ _, _) => unify_potential_flex t u |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
514 |
| (_, _ $ _) => unify_potential_flex u t |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
515 |
| (Var _, _) => add_terms (t, u) |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
516 |
| (_, Var _) => add_terms (u, t) |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
517 |
| _ => I |
42344
4a58173ffb99
"unify_first_prem_with_concl" (cf. 9ceb585c097a) sometimes throws an exception, but it is very rarely needed -- catch the exception for now
blanchet
parents:
42342
diff
changeset
|
518 |
val t_inst = |
4a58173ffb99
"unify_first_prem_with_concl" (cf. 9ceb585c097a) sometimes throws an exception, but it is very rarely needed -- catch the exception for now
blanchet
parents:
42342
diff
changeset
|
519 |
[] |> try (unify_terms (prem, concl) #> map (pairself (cterm_of thy))) |
4a58173ffb99
"unify_first_prem_with_concl" (cf. 9ceb585c097a) sometimes throws an exception, but it is very rarely needed -- catch the exception for now
blanchet
parents:
42342
diff
changeset
|
520 |
|> the_default [] (* FIXME *) |
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
521 |
in th |> cterm_instantiate t_inst end |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
522 |
|
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
523 |
val copy_prem = @{lemma "P ==> (P ==> P ==> Q) ==> Q" by fast} |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
524 |
|
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
525 |
fun copy_prems_tac [] ns i = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
526 |
if forall (curry (op =) 1) ns then all_tac else copy_prems_tac (rev ns) [] i |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
527 |
| copy_prems_tac (1 :: ms) ns i = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
528 |
rotate_tac 1 i THEN copy_prems_tac ms (1 :: ns) i |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
529 |
| copy_prems_tac (m :: ms) ns i = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
530 |
etac copy_prem i THEN copy_prems_tac ms (m div 2 :: (m + 1) div 2 :: ns) i |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
531 |
|
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
532 |
(* Metis generates variables of the form _nnn. *) |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
533 |
val is_metis_fresh_variable = String.isPrefix "_" |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
534 |
|
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
535 |
fun instantiate_forall_tac thy t i st = |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
536 |
let |
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
537 |
val params = Logic.strip_params (Logic.get_goal (prop_of st) i) |> rev |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
538 |
fun repair (t as (Var ((s, _), _))) = |
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
539 |
(case find_index (fn (s', _) => s' = s) params of |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
540 |
~1 => t |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
541 |
| j => Bound j) |
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
542 |
| repair (t $ u) = |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
543 |
(case (repair t, repair u) of |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
544 |
(t as Bound j, u as Bound k) => |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
545 |
(* This is a rather subtle trick to repair the discrepancy between |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
546 |
the fully skolemized term that MESON gives us (where existentials |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
547 |
were pulled out) and the reality. *) |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
548 |
if k > j then t else t $ u |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
549 |
| (t, u) => t $ u) |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
550 |
| repair t = t |
44241 | 551 |
val t' = t |> repair |> fold (absdummy o snd) params |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
552 |
fun do_instantiate th = |
42270 | 553 |
case Term.add_vars (prop_of th) [] |
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
554 |
|> filter_out ((Meson_Clausify.is_zapped_var_name orf |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
555 |
is_metis_fresh_variable) o fst o fst) of |
42270 | 556 |
[] => th |
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
557 |
| [var as (_, T)] => |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
558 |
let |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
559 |
val var_binder_Ts = T |> binder_types |> take (length params) |> rev |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
560 |
val var_body_T = T |> funpow (length params) range_type |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
561 |
val tyenv = |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
562 |
Vartab.empty |> Type.raw_unifys (fastype_of t :: map snd params, |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
563 |
var_body_T :: var_binder_Ts) |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
564 |
val env = |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
565 |
Envir.Envir {maxidx = Vartab.fold (Integer.max o snd o fst) tyenv 0, |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
566 |
tenv = Vartab.empty, tyenv = tyenv} |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
567 |
val ty_inst = |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
568 |
Vartab.fold (fn (x, (S, T)) => |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
569 |
cons (pairself (ctyp_of thy) (TVar (x, S), T))) |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
570 |
tyenv [] |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
571 |
val t_inst = |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
572 |
[pairself (cterm_of thy o Envir.norm_term env) (Var var, t')] |
43333
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents:
43330
diff
changeset
|
573 |
in th |> Drule.instantiate_normalize (ty_inst, t_inst) end |
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
574 |
| _ => raise Fail "expected a single non-zapped, non-Metis Var" |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
575 |
in |
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
576 |
(DETERM (etac @{thm allE} i THEN rotate_tac ~1 i) |
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
577 |
THEN PRIMITIVE do_instantiate) st |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
578 |
end |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
579 |
|
41135 | 580 |
fun fix_exists_tac t = |
43162
9a8acc5adfa3
added Metis examples to test the new type encodings
blanchet
parents:
43159
diff
changeset
|
581 |
etac @{thm exE} THEN' rename_tac [t |> dest_Var |> fst |> fst] |
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
582 |
|
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
583 |
fun release_quantifier_tac thy (skolem, t) = |
41135 | 584 |
(if skolem then fix_exists_tac else instantiate_forall_tac thy) t |
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
585 |
|
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
586 |
fun release_clusters_tac _ _ _ [] = K all_tac |
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
587 |
| release_clusters_tac thy ax_counts substs |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
588 |
((ax_no, cluster_no) :: clusters) = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
589 |
let |
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
590 |
val cluster_of_var = |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
591 |
Meson_Clausify.cluster_of_zapped_var_name o fst o fst o dest_Var |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
592 |
fun in_right_cluster ((_, (cluster_no', _)), _) = cluster_no' = cluster_no |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
593 |
val cluster_substs = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
594 |
substs |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
595 |
|> map_filter (fn (ax_no', (_, (_, tsubst))) => |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
596 |
if ax_no' = ax_no then |
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
597 |
tsubst |> map (apfst cluster_of_var) |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
598 |
|> filter (in_right_cluster o fst) |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
599 |
|> map (apfst snd) |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
600 |
|> SOME |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
601 |
else |
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
602 |
NONE) |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
603 |
fun do_cluster_subst cluster_subst = |
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40259
diff
changeset
|
604 |
map (release_quantifier_tac thy) cluster_subst @ [rotate_tac 1] |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
605 |
val first_prem = find_index (fn (ax_no', _) => ax_no' = ax_no) substs |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
606 |
in |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
607 |
rotate_tac first_prem |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
608 |
THEN' (EVERY' (maps do_cluster_subst cluster_substs)) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
609 |
THEN' rotate_tac (~ first_prem - length cluster_substs) |
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
610 |
THEN' release_clusters_tac thy ax_counts substs clusters |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
611 |
end |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
612 |
|
40264
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
blanchet
parents:
40261
diff
changeset
|
613 |
fun cluster_key ((ax_no, (cluster_no, index_no)), skolem) = |
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
blanchet
parents:
40261
diff
changeset
|
614 |
(ax_no, (cluster_no, (skolem, index_no))) |
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
blanchet
parents:
40261
diff
changeset
|
615 |
fun cluster_ord p = |
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
blanchet
parents:
40261
diff
changeset
|
616 |
prod_ord int_ord (prod_ord int_ord (prod_ord bool_ord int_ord)) |
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
blanchet
parents:
40261
diff
changeset
|
617 |
(pairself cluster_key p) |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
618 |
|
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
619 |
val tysubst_ord = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
620 |
list_ord (prod_ord Term_Ord.fast_indexname_ord |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
621 |
(prod_ord Term_Ord.sort_ord Term_Ord.typ_ord)) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
622 |
|
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
623 |
structure Int_Tysubst_Table = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
624 |
Table(type key = int * (indexname * (sort * typ)) list |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
625 |
val ord = prod_ord int_ord tysubst_ord) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
626 |
|
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
627 |
structure Int_Pair_Graph = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
628 |
Graph(type key = int * int val ord = prod_ord int_ord int_ord) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
629 |
|
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
630 |
fun shuffle_key (((axiom_no, (_, index_no)), _), _) = (axiom_no, index_no) |
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
631 |
fun shuffle_ord p = prod_ord int_ord int_ord (pairself shuffle_key p) |
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
632 |
|
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
633 |
(* Attempts to derive the theorem "False" from a theorem of the form |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
634 |
"P1 ==> ... ==> Pn ==> False", where the "Pi"s are to be discharged using the |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
635 |
specified axioms. The axioms have leading "All" and "Ex" quantifiers, which |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
636 |
must be eliminated first. *) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
637 |
fun discharge_skolem_premises ctxt axioms prems_imp_false = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
638 |
if prop_of prems_imp_false aconv @{prop False} then |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
639 |
prems_imp_false |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
640 |
else |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
641 |
let |
42361 | 642 |
val thy = Proof_Context.theory_of ctxt |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
643 |
fun match_term p = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
644 |
let |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
645 |
val (tyenv, tenv) = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
646 |
Pattern.first_order_match thy p (Vartab.empty, Vartab.empty) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
647 |
val tsubst = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
648 |
tenv |> Vartab.dest |
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
649 |
|> filter (Meson_Clausify.is_zapped_var_name o fst o fst) |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
650 |
|> sort (cluster_ord |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
651 |
o pairself (Meson_Clausify.cluster_of_zapped_var_name |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
652 |
o fst o fst)) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
653 |
|> map (Meson.term_pair_of |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
654 |
#> pairself (Envir.subst_term_types tyenv)) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
655 |
val tysubst = tyenv |> Vartab.dest |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
656 |
in (tysubst, tsubst) end |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
657 |
fun subst_info_for_prem subgoal_no prem = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
658 |
case prem of |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
659 |
_ $ (Const (@{const_name Meson.skolem}, _) $ (_ $ t $ num)) => |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
660 |
let val ax_no = HOLogic.dest_nat num in |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
661 |
(ax_no, (subgoal_no, |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
662 |
match_term (nth axioms ax_no |> the |> snd, t))) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
663 |
end |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
664 |
| _ => raise TERM ("discharge_skolem_premises: Malformed premise", |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
665 |
[prem]) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
666 |
fun cluster_of_var_name skolem s = |
42098
f978caf60bbe
more robust handling of variables in new Skolemizer
blanchet
parents:
41491
diff
changeset
|
667 |
case try Meson_Clausify.cluster_of_zapped_var_name s of |
f978caf60bbe
more robust handling of variables in new Skolemizer
blanchet
parents:
41491
diff
changeset
|
668 |
NONE => NONE |
f978caf60bbe
more robust handling of variables in new Skolemizer
blanchet
parents:
41491
diff
changeset
|
669 |
| SOME ((ax_no, (cluster_no, _)), skolem') => |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
670 |
if skolem' = skolem andalso cluster_no > 0 then |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
671 |
SOME (ax_no, cluster_no) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
672 |
else |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
673 |
NONE |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
674 |
fun clusters_in_term skolem t = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
675 |
Term.add_var_names t [] |> map_filter (cluster_of_var_name skolem o fst) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
676 |
fun deps_for_term_subst (var, t) = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
677 |
case clusters_in_term false var of |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
678 |
[] => NONE |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
679 |
| [(ax_no, cluster_no)] => |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
680 |
SOME ((ax_no, cluster_no), |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
681 |
clusters_in_term true t |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
682 |
|> cluster_no > 1 ? cons (ax_no, cluster_no - 1)) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
683 |
| _ => raise TERM ("discharge_skolem_premises: Expected Var", [var]) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
684 |
val prems = Logic.strip_imp_prems (prop_of prems_imp_false) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
685 |
val substs = prems |> map2 subst_info_for_prem (1 upto length prems) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
686 |
|> sort (int_ord o pairself fst) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
687 |
val depss = maps (map_filter deps_for_term_subst o snd o snd o snd) substs |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
688 |
val clusters = maps (op ::) depss |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
689 |
val ordered_clusters = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
690 |
Int_Pair_Graph.empty |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
691 |
|> fold Int_Pair_Graph.default_node (map (rpair ()) clusters) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
692 |
|> fold Int_Pair_Graph.add_deps_acyclic depss |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
693 |
|> Int_Pair_Graph.topological_order |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
694 |
handle Int_Pair_Graph.CYCLES _ => |
40158 | 695 |
error "Cannot replay Metis proof in Isabelle without \ |
696 |
\\"Hilbert_Choice\"." |
|
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
697 |
val ax_counts = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
698 |
Int_Tysubst_Table.empty |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
699 |
|> fold (fn (ax_no, (_, (tysubst, _))) => |
43262 | 700 |
Int_Tysubst_Table.map_default ((ax_no, tysubst), 0) |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
701 |
(Integer.add 1)) substs |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
702 |
|> Int_Tysubst_Table.dest |
42339
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
703 |
val needed_axiom_props = |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
704 |
0 upto length axioms - 1 ~~ axioms |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
705 |
|> map_filter (fn (_, NONE) => NONE |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
706 |
| (ax_no, SOME (_, t)) => |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
707 |
if exists (fn ((ax_no', _), n) => |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
708 |
ax_no' = ax_no andalso n > 0) |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
709 |
ax_counts then |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
710 |
SOME t |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
711 |
else |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
712 |
NONE) |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
713 |
val outer_param_names = |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
714 |
[] |> fold Term.add_var_names needed_axiom_props |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
715 |
|> filter (Meson_Clausify.is_zapped_var_name o fst) |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
716 |
|> map (`(Meson_Clausify.cluster_of_zapped_var_name o fst)) |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
717 |
|> filter (fn (((_, (cluster_no, _)), skolem), _) => |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
718 |
cluster_no = 0 andalso skolem) |
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
719 |
|> sort shuffle_ord |> map (fst o snd) |
42270 | 720 |
(* for debugging only: |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
721 |
fun string_for_subst_info (ax_no, (subgoal_no, (tysubst, tsubst))) = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
722 |
"ax: " ^ string_of_int ax_no ^ "; asm: " ^ string_of_int subgoal_no ^ |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
723 |
"; tysubst: " ^ PolyML.makestring tysubst ^ "; tsubst: {" ^ |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
724 |
commas (map ((fn (s, t) => s ^ " |-> " ^ t) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
725 |
o pairself (Syntax.string_of_term ctxt)) tsubst) ^ "}" |
40264
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
blanchet
parents:
40261
diff
changeset
|
726 |
val _ = tracing ("ORDERED CLUSTERS: " ^ PolyML.makestring ordered_clusters) |
b91e2e16d994
fixed order of quantifier instantiation in new Skolemizer
blanchet
parents:
40261
diff
changeset
|
727 |
val _ = tracing ("AXIOM COUNTS: " ^ PolyML.makestring ax_counts) |
42339
0e5d1e5e1177
use the list of actually used axioms to (correctly) precompute the "outer params", not all axioms
blanchet
parents:
42337
diff
changeset
|
728 |
val _ = tracing ("OUTER PARAMS: " ^ PolyML.makestring outer_param_names) |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
729 |
val _ = tracing ("SUBSTS (" ^ string_of_int (length substs) ^ "):\n" ^ |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
730 |
cat_lines (map string_for_subst_info substs)) |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
731 |
*) |
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
732 |
fun cut_and_ex_tac axiom = |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
733 |
cut_rules_tac [axiom] 1 |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
734 |
THEN TRY (REPEAT_ALL_NEW (etac @{thm exE}) 1) |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
735 |
fun rotation_for_subgoal i = |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
736 |
find_index (fn (_, (subgoal_no, _)) => subgoal_no = i) substs |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
737 |
in |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
738 |
Goal.prove ctxt [] [] @{prop False} |
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
739 |
(K (DETERM (EVERY (map (cut_and_ex_tac o fst o the o nth axioms o fst |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
740 |
o fst) ax_counts) |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
741 |
THEN rename_tac outer_param_names 1 |
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
742 |
THEN copy_prems_tac (map snd ax_counts) [] 1) |
40258
2c0d8fe36c21
make handling of parameters more robust, by querying the goal
blanchet
parents:
40221
diff
changeset
|
743 |
THEN release_clusters_tac thy ax_counts substs ordered_clusters 1 |
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
744 |
THEN match_tac [prems_imp_false] 1 |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
745 |
THEN ALLGOALS (fn i => |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
746 |
rtac @{thm Meson.skolem_COMBK_I} i |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
747 |
THEN rotate_tac (rotation_for_subgoal i) i |
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42341
diff
changeset
|
748 |
THEN PRIMITIVE (unify_first_prem_with_concl thy i) |
42271
7d08265f181d
further development of new Skolemizer -- make sure constructed terms have correct types and fixed a few bugs where the goal was out of sync with what we had in mind
blanchet
parents:
42270
diff
changeset
|
749 |
THEN assume_tac i |
42270 | 750 |
THEN flexflex_tac))) |
40158 | 751 |
handle ERROR _ => |
752 |
error ("Cannot replay Metis proof in Isabelle:\n\ |
|
753 |
\Error when discharging Skolem assumptions.") |
|
39964
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
754 |
end |
8ca95d819c7c
move code from "Metis_Tactics" to "Metis_Reconstruct"
blanchet
parents:
39958
diff
changeset
|
755 |
|
39495
bb4fb9ffe2d1
added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff
changeset
|
756 |
end; |