| author | bulwahn |
| Wed, 10 Oct 2012 10:41:18 +0200 | |
| changeset 49762 | b5e355c41de3 |
| parent 49740 | 6f02b893dd87 |
| child 49881 | d9d73ebf9274 |
| permissions | -rw-r--r-- |
| 46320 | 1 |
(* Title: HOL/Tools/ATP/atp_proof_reconstruct.ML |
| 38027 | 2 |
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory |
3 |
Author: Claire Quigley, Cambridge University Computer Laboratory |
|
| 36392 | 4 |
Author: Jasmin Blanchette, TU Muenchen |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
5 |
|
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43062
diff
changeset
|
6 |
Proof reconstruction from ATP proofs. |
| 33310 | 7 |
*) |
8 |
||
| 46320 | 9 |
signature ATP_PROOF_RECONSTRUCT = |
|
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
10 |
sig |
| 43678 | 11 |
type ('a, 'b) ho_term = ('a, 'b) ATP_Problem.ho_term
|
| 48135 | 12 |
type ('a, 'b, 'c, 'd) formula = ('a, 'b, 'c, 'd) ATP_Problem.formula
|
|
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
13 |
type 'a proof = 'a ATP_Proof.proof |
| 46340 | 14 |
type stature = ATP_Problem_Generate.stature |
| 43033 | 15 |
|
16 |
datatype reconstructor = |
|
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
17 |
Metis of string * string | |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
18 |
SMT |
| 43033 | 19 |
|
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43039
diff
changeset
|
20 |
datatype play = |
|
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43039
diff
changeset
|
21 |
Played of reconstructor * Time.time | |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
22 |
Trust_Playable of reconstructor * Time.time option | |
| 43166 | 23 |
Failed_to_Play of reconstructor |
| 43033 | 24 |
|
|
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
25 |
type minimize_command = string list -> string |
| 43033 | 26 |
type one_line_params = |
| 46340 | 27 |
play * string * (string * stature) list * minimize_command * int * int |
|
38818
61cf050f8b2e
improve SPASS hack, when a clause comes from several facts
blanchet
parents:
38752
diff
changeset
|
28 |
type isar_params = |
| 46340 | 29 |
bool * int * string Symtab.table * (string * stature) list vector |
| 45551 | 30 |
* int Symtab.table * string proof * thm |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
31 |
|
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
32 |
val metisN : string |
| 45520 | 33 |
val smtN : string |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
34 |
val full_typesN : string |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
35 |
val partial_typesN : string |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
36 |
val no_typesN : string |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
37 |
val really_full_type_enc : string |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
38 |
val full_type_enc : string |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
39 |
val partial_type_enc : string |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
40 |
val no_type_enc : string |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
41 |
val full_type_encs : string list |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
42 |
val partial_type_encs : string list |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
43 |
val metis_default_lam_trans : string |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
44 |
val metis_call : string -> string -> string |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
45 |
val string_for_reconstructor : reconstructor -> string |
|
42451
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42449
diff
changeset
|
46 |
val used_facts_in_atp_proof : |
| 46340 | 47 |
Proof.context -> (string * stature) list vector -> string proof |
48 |
-> (string * stature) list |
|
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
49 |
val lam_trans_from_atp_proof : string proof -> string -> string |
| 45590 | 50 |
val is_typed_helper_used_in_atp_proof : string proof -> bool |
|
42876
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
blanchet
parents:
42761
diff
changeset
|
51 |
val used_facts_in_unsound_atp_proof : |
| 46340 | 52 |
Proof.context -> (string * stature) list vector -> 'a proof |
| 45551 | 53 |
-> string list option |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
54 |
val unalias_type_enc : string -> string list |
| 48799 | 55 |
val one_line_proof_text : int -> one_line_params -> string |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
56 |
val make_tvar : string -> typ |
|
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
57 |
val make_tfree : Proof.context -> string -> typ |
| 43094 | 58 |
val term_from_atp : |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
59 |
Proof.context -> bool -> int Symtab.table -> typ option |
|
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
60 |
-> (string, string) ho_term -> term |
| 43127 | 61 |
val prop_from_atp : |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
62 |
Proof.context -> bool -> int Symtab.table |
| 48135 | 63 |
-> (string, string, (string, string) ho_term, string) formula -> term |
|
42968
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents:
42966
diff
changeset
|
64 |
val isar_proof_text : |
|
43062
2834548a7a48
nicer failure message when one-line proof reconstruction fails
blanchet
parents:
43051
diff
changeset
|
65 |
Proof.context -> bool -> isar_params -> one_line_params -> string |
|
42968
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents:
42966
diff
changeset
|
66 |
val proof_text : |
| 48799 | 67 |
Proof.context -> bool -> isar_params -> int -> one_line_params -> string |
|
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
68 |
end; |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
69 |
|
| 46320 | 70 |
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT = |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
71 |
struct |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
72 |
|
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43062
diff
changeset
|
73 |
open ATP_Util |
| 38028 | 74 |
open ATP_Problem |
| 39452 | 75 |
open ATP_Proof |
| 46320 | 76 |
open ATP_Problem_Generate |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
77 |
|
| 46320 | 78 |
structure String_Redirect = ATP_Proof_Redirect( |
| 45887 | 79 |
type key = step_name |
80 |
val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s') |
|
81 |
val string_of = fst) |
|
82 |
||
| 45882 | 83 |
open String_Redirect |
84 |
||
| 43033 | 85 |
datatype reconstructor = |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
86 |
Metis of string * string | |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
87 |
SMT |
| 43033 | 88 |
|
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43039
diff
changeset
|
89 |
datatype play = |
|
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43039
diff
changeset
|
90 |
Played of reconstructor * Time.time | |
| 43033 | 91 |
Trust_Playable of reconstructor * Time.time option | |
| 43166 | 92 |
Failed_to_Play of reconstructor |
| 43033 | 93 |
|
|
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
94 |
type minimize_command = string list -> string |
| 43033 | 95 |
type one_line_params = |
| 46340 | 96 |
play * string * (string * stature) list * minimize_command * int * int |
|
38818
61cf050f8b2e
improve SPASS hack, when a clause comes from several facts
blanchet
parents:
38752
diff
changeset
|
97 |
type isar_params = |
| 46340 | 98 |
bool * int * string Symtab.table * (string * stature) list vector |
| 45551 | 99 |
* int Symtab.table * string proof * thm |
|
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
100 |
|
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
101 |
val metisN = "metis" |
| 45520 | 102 |
val smtN = "smt" |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
103 |
|
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
104 |
val full_typesN = "full_types" |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
105 |
val partial_typesN = "partial_types" |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
106 |
val no_typesN = "no_types" |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
107 |
|
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
108 |
val really_full_type_enc = "mono_tags" |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
109 |
val full_type_enc = "poly_guards_query" |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
110 |
val partial_type_enc = "poly_args" |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
111 |
val no_type_enc = "erased" |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
112 |
|
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
113 |
val full_type_encs = [full_type_enc, really_full_type_enc] |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
114 |
val partial_type_encs = partial_type_enc :: full_type_encs |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
115 |
|
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
116 |
val type_enc_aliases = |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
117 |
[(full_typesN, full_type_encs), |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
118 |
(partial_typesN, partial_type_encs), |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
119 |
(no_typesN, [no_type_enc])] |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
120 |
|
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
121 |
fun unalias_type_enc s = |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
122 |
AList.lookup (op =) type_enc_aliases s |> the_default [s] |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
123 |
|
| 46365 | 124 |
val metis_default_lam_trans = combsN |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
125 |
|
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
126 |
fun metis_call type_enc lam_trans = |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
127 |
let |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
128 |
val type_enc = |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
129 |
case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
130 |
type_enc of |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
131 |
[alias] => alias |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
132 |
| _ => type_enc |
| 45522 | 133 |
val opts = [] |> type_enc <> partial_typesN ? cons type_enc |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
134 |
|> lam_trans <> metis_default_lam_trans ? cons lam_trans |
| 45520 | 135 |
in metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") end
|
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
136 |
|
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
137 |
fun string_for_reconstructor (Metis (type_enc, lam_trans)) = |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
138 |
metis_call type_enc lam_trans |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
139 |
| string_for_reconstructor SMT = smtN |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
140 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
141 |
fun find_first_in_list_vector vec key = |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
142 |
Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
143 |
| (_, value) => value) NONE vec |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
144 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
145 |
val unprefix_fact_number = space_implode "_" o tl o space_explode "_" |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
146 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
147 |
fun resolve_one_named_fact fact_names s = |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
148 |
case try (unprefix fact_prefix) s of |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
149 |
SOME s' => |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
150 |
let val s' = s' |> unprefix_fact_number |> unascii_of in |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
151 |
s' |> find_first_in_list_vector fact_names |> Option.map (pair s') |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
152 |
end |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
153 |
| NONE => NONE |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
154 |
fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names) |
| 45579 | 155 |
fun is_fact fact_names = not o null o resolve_fact fact_names |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
156 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
157 |
fun resolve_one_named_conjecture s = |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
158 |
case try (unprefix conjecture_prefix) s of |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
159 |
SOME s' => Int.fromString s' |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
160 |
| NONE => NONE |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
161 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
162 |
val resolve_conjecture = map_filter resolve_one_named_conjecture |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
163 |
val is_conjecture = not o null o resolve_conjecture |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
164 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
165 |
fun is_axiom_used_in_proof pred = |
| 47774 | 166 |
exists (fn Inference_Step ((_, ss), _, _, []) => exists pred ss | _ => false) |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
167 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
168 |
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix) |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
169 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
170 |
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
171 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
172 |
(* overapproximation (good enough) *) |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
173 |
fun is_lam_lifted s = |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
174 |
String.isPrefix fact_prefix s andalso |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
175 |
String.isSubstring ascii_of_lam_fact_prefix s |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
176 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
177 |
fun lam_trans_from_atp_proof atp_proof default = |
| 46365 | 178 |
case (is_axiom_used_in_proof is_combinator_def atp_proof, |
179 |
is_axiom_used_in_proof is_lam_lifted atp_proof) of |
|
180 |
(false, false) => default |
|
181 |
| (false, true) => liftingN |
|
| 46367 | 182 |
(* | (true, true) => combs_and_liftingN -- not supported by "metis" *) |
183 |
| (true, _) => combsN |
|
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
184 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
185 |
val is_typed_helper_name = |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
186 |
String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix |
| 45590 | 187 |
fun is_typed_helper_used_in_atp_proof atp_proof = |
188 |
is_axiom_used_in_proof is_typed_helper_name atp_proof |
|
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
189 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
190 |
val leo2_ext = "extcnf_equal_neg" |
|
47974
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47973
diff
changeset
|
191 |
val leo2_unfold_def = "unfold_def" |
|
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47973
diff
changeset
|
192 |
|
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
193 |
val isa_ext = Thm.get_name_hint @{thm ext}
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
194 |
val isa_short_ext = Long_Name.base_name isa_ext |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
195 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
196 |
fun ext_name ctxt = |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
197 |
if Thm.eq_thm_prop (@{thm ext},
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
198 |
singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
199 |
isa_short_ext |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
200 |
else |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
201 |
isa_ext |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
202 |
|
|
48438
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48135
diff
changeset
|
203 |
fun add_non_rec_defs fact_names accum = |
|
48085
ff5e900d7b1a
avoid dumping definitions several times in LEO-II proofs
blanchet
parents:
47974
diff
changeset
|
204 |
Vector.foldl |
|
ff5e900d7b1a
avoid dumping definitions several times in LEO-II proofs
blanchet
parents:
47974
diff
changeset
|
205 |
(fn (facts, facts') => |
|
48438
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48135
diff
changeset
|
206 |
union (op =) |
|
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48135
diff
changeset
|
207 |
(filter (fn (_, (_, status)) => status = Non_Rec_Def) facts) |
|
48085
ff5e900d7b1a
avoid dumping definitions several times in LEO-II proofs
blanchet
parents:
47974
diff
changeset
|
208 |
facts') |
|
ff5e900d7b1a
avoid dumping definitions several times in LEO-II proofs
blanchet
parents:
47974
diff
changeset
|
209 |
accum fact_names |
|
47974
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47973
diff
changeset
|
210 |
|
| 47947 | 211 |
fun add_fact ctxt fact_names (Inference_Step ((_, ss), _, rule, deps)) = |
|
47973
9af7e5caf16f
augment Satallax unsat cores with all definitions
blanchet
parents:
47947
diff
changeset
|
212 |
(if rule = leo2_ext then |
| 47947 | 213 |
insert (op =) (ext_name ctxt, (Global, General)) |
|
47974
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47973
diff
changeset
|
214 |
else if rule = leo2_unfold_def then |
|
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47973
diff
changeset
|
215 |
(* LEO 1.3.3 does not record definitions properly, leading to missing |
|
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47973
diff
changeset
|
216 |
dependencies in the TSTP proof. Remove the next line once this is |
|
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47973
diff
changeset
|
217 |
fixed. *) |
|
48438
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48135
diff
changeset
|
218 |
add_non_rec_defs fact_names |
| 48539 | 219 |
else if rule = satallax_coreN then |
|
47973
9af7e5caf16f
augment Satallax unsat cores with all definitions
blanchet
parents:
47947
diff
changeset
|
220 |
(fn [] => |
|
9af7e5caf16f
augment Satallax unsat cores with all definitions
blanchet
parents:
47947
diff
changeset
|
221 |
(* Satallax doesn't include definitions in its unsatisfiable cores, |
|
9af7e5caf16f
augment Satallax unsat cores with all definitions
blanchet
parents:
47947
diff
changeset
|
222 |
so we assume the worst and include them all here. *) |
|
48438
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48135
diff
changeset
|
223 |
[(ext_name ctxt, (Global, General))] |> add_non_rec_defs fact_names |
|
47973
9af7e5caf16f
augment Satallax unsat cores with all definitions
blanchet
parents:
47947
diff
changeset
|
224 |
| facts => facts) |
| 47947 | 225 |
else |
226 |
I) |
|
227 |
#> (if null deps then union (op =) (resolve_fact fact_names ss) |
|
228 |
else I) |
|
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
229 |
| add_fact _ _ _ = I |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
230 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
231 |
fun used_facts_in_atp_proof ctxt fact_names atp_proof = |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
232 |
if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
233 |
else fold (add_fact ctxt fact_names) atp_proof [] |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
234 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
235 |
fun used_facts_in_unsound_atp_proof _ _ [] = NONE |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
236 |
| used_facts_in_unsound_atp_proof ctxt fact_names atp_proof = |
| 46427 | 237 |
let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in |
| 46340 | 238 |
if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
239 |
not (is_axiom_used_in_proof (is_conjecture o single) atp_proof) then |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
240 |
SOME (map fst used_facts) |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
241 |
else |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
242 |
NONE |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
243 |
end |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
244 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
245 |
|
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
246 |
(** Soft-core proof reconstruction: one-liners **) |
| 43033 | 247 |
|
|
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
248 |
fun string_for_label (s, num) = s ^ string_of_int num |
|
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
249 |
|
| 43033 | 250 |
fun show_time NONE = "" |
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
251 |
| show_time (SOME ext_time) = " (" ^ string_from_ext_time ext_time ^ ")"
|
| 43033 | 252 |
|
| 48799 | 253 |
fun unusing_chained_facts _ 0 = "" |
254 |
| unusing_chained_facts used_chaineds num_chained = |
|
255 |
if length used_chaineds = num_chained then "" |
|
256 |
else if null used_chaineds then "(* using no facts *) " |
|
257 |
else "(* using only " ^ space_implode " " used_chaineds ^ " *) " |
|
258 |
||
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
259 |
fun apply_on_subgoal _ 1 = "by " |
|
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
260 |
| apply_on_subgoal 1 _ = "apply " |
|
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
261 |
| apply_on_subgoal i n = |
|
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
262 |
"prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n |
| 48799 | 263 |
|
| 44002 | 264 |
fun command_call name [] = |
265 |
name |> not (Lexicon.is_identifier name) ? enclose "(" ")"
|
|
|
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
266 |
| command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
|
| 48799 | 267 |
|
| 43033 | 268 |
fun try_command_line banner time command = |
| 45666 | 269 |
banner ^ ": " ^ Markup.markup Isabelle_Markup.sendback command ^ show_time time ^ "." |
| 48799 | 270 |
|
|
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
271 |
fun using_labels [] = "" |
|
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
272 |
| using_labels ls = |
|
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
273 |
"using " ^ space_implode " " (map string_for_label ls) ^ " " |
| 48799 | 274 |
|
275 |
fun reconstructor_command reconstr i n used_chaineds num_chained (ls, ss) = |
|
276 |
unusing_chained_facts used_chaineds num_chained ^ |
|
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
277 |
using_labels ls ^ apply_on_subgoal i n ^ |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
278 |
command_call (string_for_reconstructor reconstr) ss |
| 48799 | 279 |
|
|
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
280 |
fun minimize_line _ [] = "" |
|
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
281 |
| minimize_line minimize_command ss = |
|
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
282 |
case minimize_command ss of |
|
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42361
diff
changeset
|
283 |
"" => "" |
| 46340 | 284 |
| command => |
285 |
"\nTo minimize: " ^ Markup.markup Isabelle_Markup.sendback command ^ "." |
|
|
39453
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
286 |
|
| 46342 | 287 |
fun split_used_facts facts = |
288 |
facts |> List.partition (fn (_, (sc, _)) => sc = Chained) |
|
289 |
|> pairself (sort_distinct (string_ord o pairself fst)) |
|
|
39453
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
290 |
|
| 48799 | 291 |
fun one_line_proof_text num_chained |
292 |
(preplay, banner, used_facts, minimize_command, subgoal, |
|
293 |
subgoal_count) = |
|
|
39453
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
294 |
let |
| 43033 | 295 |
val (chained, extra) = split_used_facts used_facts |
| 45520 | 296 |
val (failed, reconstr, ext_time) = |
| 43033 | 297 |
case preplay of |
| 45520 | 298 |
Played (reconstr, time) => (false, reconstr, (SOME (false, time))) |
299 |
| Trust_Playable (reconstr, time) => |
|
300 |
(false, reconstr, |
|
| 43033 | 301 |
case time of |
302 |
NONE => NONE |
|
303 |
| SOME time => |
|
304 |
if time = Time.zeroTime then NONE else SOME (true, time)) |
|
| 45520 | 305 |
| Failed_to_Play reconstr => (true, reconstr, NONE) |
| 43033 | 306 |
val try_line = |
| 43166 | 307 |
([], map fst extra) |
| 48799 | 308 |
|> reconstructor_command reconstr subgoal subgoal_count (map fst chained) |
309 |
num_chained |
|
| 47147 | 310 |
|> (if failed then |
311 |
enclose "One-line proof reconstruction failed: " |
|
312 |
".\n(Invoking \"sledgehammer\" with \"[strict]\" might \ |
|
313 |
\solve this.)" |
|
314 |
else |
|
315 |
try_command_line banner ext_time) |
|
| 43033 | 316 |
in try_line ^ minimize_line minimize_command (map fst (extra @ chained)) end |
|
39453
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
317 |
|
|
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
318 |
(** Hard-core proof reconstruction: structured Isar proofs **) |
|
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
319 |
|
| 39425 | 320 |
fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t |
321 |
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t |
|
322 |
||
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
323 |
fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
|
|
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
324 |
fun make_tfree ctxt w = |
|
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
325 |
let val ww = "'" ^ w in |
|
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
326 |
TFree (ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1))) |
|
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
327 |
end |
|
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
328 |
|
|
39453
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
329 |
val indent_size = 2 |
|
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
330 |
val no_label = ("", ~1)
|
|
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
331 |
|
| 45882 | 332 |
val raw_prefix = "x" |
333 |
val assum_prefix = "a" |
|
334 |
val have_prefix = "f" |
|
|
39370
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents:
39368
diff
changeset
|
335 |
|
| 45552 | 336 |
fun raw_label_for_name (num, ss) = |
337 |
case resolve_conjecture ss of |
|
|
39453
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
338 |
[j] => (conjecture_prefix, j) |
|
47920
a5c2386518e2
ensure consistent naming of Waldmeister proof steps, so that they are not cleaned away by "clean_up_atp_proof_dependencies"
blanchet
parents:
47918
diff
changeset
|
339 |
| _ => (raw_prefix ^ ascii_of num, 0) |
|
39453
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents:
39452
diff
changeset
|
340 |
|
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
341 |
(**** INTERPRETATION OF TSTP SYNTAX TREES ****) |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
342 |
|
| 43678 | 343 |
exception HO_TERM of (string, string) ho_term list |
| 48135 | 344 |
exception FORMULA of |
345 |
(string, string, (string, string) ho_term, string) formula list |
|
| 37991 | 346 |
exception SAME of unit |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
347 |
|
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
348 |
(* Type variables are given the basic sort "HOL.type". Some will later be |
| 37991 | 349 |
constrained by information from type literals, or by type inference. *) |
|
48132
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents:
48085
diff
changeset
|
350 |
fun typ_from_atp ctxt (u as ATerm ((a, _), us)) = |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
351 |
let val Ts = map (typ_from_atp ctxt) us in |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
352 |
case unprefix_and_unascii type_const_prefix a of |
| 37991 | 353 |
SOME b => Type (invert_const b, Ts) |
354 |
| NONE => |
|
355 |
if not (null us) then |
|
| 43678 | 356 |
raise HO_TERM [u] (* only "tconst"s have type arguments *) |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
357 |
else case unprefix_and_unascii tfree_prefix a of |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
358 |
SOME b => make_tfree ctxt b |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
359 |
| NONE => |
|
43302
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents:
43296
diff
changeset
|
360 |
(* Could be an Isabelle variable or a variable from the ATP, say "X1" |
|
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents:
43296
diff
changeset
|
361 |
or "_5018". Sometimes variables from the ATP are indistinguishable |
|
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents:
43296
diff
changeset
|
362 |
from Isabelle variables, which forces us to use a type parameter in |
|
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents:
43296
diff
changeset
|
363 |
all cases. *) |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
364 |
(a |> perhaps (unprefix_and_unascii tvar_prefix), HOLogic.typeS) |
|
43302
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents:
43296
diff
changeset
|
365 |
|> Type_Infer.param 0 |
| 37991 | 366 |
end |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
367 |
|
| 38014 | 368 |
(* Type class literal applied to a type. Returns triple of polarity, class, |
369 |
type. *) |
|
|
48132
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents:
48085
diff
changeset
|
370 |
fun type_constraint_from_term ctxt (u as ATerm ((a, _), us)) = |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
371 |
case (unprefix_and_unascii class_prefix a, map (typ_from_atp ctxt) us) of |
|
42606
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
372 |
(SOME b, [T]) => (b, T) |
| 43678 | 373 |
| _ => raise HO_TERM [u] |
| 38014 | 374 |
|
|
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
375 |
(* Accumulate type constraints in a formula: negative type literals. *) |
| 38014 | 376 |
fun add_var (key, z) = Vartab.map_default (key, []) (cons z) |
|
42606
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
377 |
fun add_type_constraint false (cl, TFree (a ,_)) = add_var ((a, ~1), cl) |
|
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
378 |
| add_type_constraint false (cl, TVar (ix, _)) = add_var (ix, cl) |
|
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
379 |
| add_type_constraint _ _ = I |
| 38014 | 380 |
|
| 43094 | 381 |
fun repair_variable_name f s = |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
382 |
let |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
383 |
fun subscript_name s n = s ^ nat_subscript n |
| 38488 | 384 |
val s = String.map f s |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
385 |
in |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
386 |
case space_explode "_" s of |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
387 |
[_] => (case take_suffix Char.isDigit (String.explode s) of |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
388 |
(cs1 as _ :: _, cs2 as _ :: _) => |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
389 |
subscript_name (String.implode cs1) |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
390 |
(the (Int.fromString (String.implode cs2))) |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
391 |
| (_, _) => s) |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
392 |
| [s1, s2] => (case Int.fromString s2 of |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
393 |
SOME n => subscript_name s1 n |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
394 |
| NONE => s) |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
395 |
| _ => s |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
396 |
end |
|
43182
649bada59658
slacker version of "fastype_of", in case a function has dummy type
blanchet
parents:
43176
diff
changeset
|
397 |
|
|
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
398 |
(* The number of type arguments of a constant, zero if it's monomorphic. For |
|
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
399 |
(instances of) Skolem pseudoconstants, this information is encoded in the |
|
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
400 |
constant name. *) |
|
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
401 |
fun num_type_args thy s = |
|
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
402 |
if String.isPrefix skolem_const_prefix s then |
|
46711
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents:
46435
diff
changeset
|
403 |
s |> Long_Name.explode |> List.last |> Int.fromString |> the |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
404 |
else if String.isPrefix lam_lifted_prefix s then |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
405 |
if String.isPrefix lam_lifted_poly_prefix s then 2 else 0 |
|
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
406 |
else |
|
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
407 |
(s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length |
|
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
408 |
|
|
43182
649bada59658
slacker version of "fastype_of", in case a function has dummy type
blanchet
parents:
43176
diff
changeset
|
409 |
fun slack_fastype_of t = fastype_of t handle TERM _ => HOLogic.typeT |
|
649bada59658
slacker version of "fastype_of", in case a function has dummy type
blanchet
parents:
43176
diff
changeset
|
410 |
|
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
411 |
(* First-order translation. No types are known for variables. "HOLogic.typeT" |
| 38014 | 412 |
should allow them to be inferred. *) |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
413 |
fun term_from_atp ctxt textual sym_tab = |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
414 |
let |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
415 |
val thy = Proof_Context.theory_of ctxt |
| 43212 | 416 |
(* For Metis, we use 1 rather than 0 because variable references in clauses |
417 |
may otherwise conflict with variable constraints in the goal. At least, |
|
418 |
type inference often fails otherwise. See also "axiom_inference" in |
|
419 |
"Metis_Reconstruct". *) |
|
| 43094 | 420 |
val var_index = if textual then 0 else 1 |
|
43131
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
421 |
fun do_term extra_ts opt_T u = |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
422 |
case u of |
|
48132
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents:
48085
diff
changeset
|
423 |
ATerm ((s, _), us) => |
| 46435 | 424 |
if String.isPrefix native_type_prefix s then |
|
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42943
diff
changeset
|
425 |
@{const True} (* ignore TPTP type information *)
|
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
426 |
else if s = tptp_equal then |
| 43093 | 427 |
let val ts = map (do_term [] NONE) us in |
| 43094 | 428 |
if textual andalso length ts = 2 andalso |
429 |
hd ts aconv List.last ts then |
|
| 39106 | 430 |
(* Vampire is keen on producing these. *) |
431 |
@{const True}
|
|
432 |
else |
|
433 |
list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
|
|
434 |
end |
|
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
435 |
else case unprefix_and_unascii const_prefix s of |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
436 |
SOME s' => |
|
42761
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
437 |
let |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
438 |
val ((s', s''), mangled_us) = |
|
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
439 |
s' |> unmangled_const |>> `invert_const |
|
42761
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
440 |
in |
|
42755
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents:
42751
diff
changeset
|
441 |
if s' = type_tag_name then |
|
42589
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42587
diff
changeset
|
442 |
case mangled_us @ us of |
|
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42587
diff
changeset
|
443 |
[typ_u, term_u] => |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
444 |
do_term extra_ts (SOME (typ_from_atp ctxt typ_u)) term_u |
| 43678 | 445 |
| _ => raise HO_TERM us |
|
42966
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
blanchet
parents:
42962
diff
changeset
|
446 |
else if s' = predicator_name then |
| 43093 | 447 |
do_term [] (SOME @{typ bool}) (hd us)
|
|
42966
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
blanchet
parents:
42962
diff
changeset
|
448 |
else if s' = app_op_name then |
|
43131
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
449 |
let val extra_t = do_term [] NONE (List.last us) in |
|
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
450 |
do_term (extra_t :: extra_ts) |
|
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
451 |
(case opt_T of |
|
43182
649bada59658
slacker version of "fastype_of", in case a function has dummy type
blanchet
parents:
43176
diff
changeset
|
452 |
SOME T => SOME (slack_fastype_of extra_t --> T) |
|
43131
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
453 |
| NONE => NONE) |
|
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
454 |
(nth us (length us - 2)) |
|
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
455 |
end |
|
44396
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44121
diff
changeset
|
456 |
else if s' = type_guard_name then |
|
42551
cd99d6d3027a
reconstruct TFF type predicates correctly for ToFoF
blanchet
parents:
42549
diff
changeset
|
457 |
@{const True} (* ignore type predicates *)
|
|
42549
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents:
42544
diff
changeset
|
458 |
else |
|
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents:
42544
diff
changeset
|
459 |
let |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
460 |
val new_skolem = String.isPrefix new_skolem_const_prefix s'' |
|
42755
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents:
42751
diff
changeset
|
461 |
val num_ty_args = |
|
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents:
42751
diff
changeset
|
462 |
length us - the_default 0 (Symtab.lookup sym_tab s) |
|
42549
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents:
42544
diff
changeset
|
463 |
val (type_us, term_us) = |
|
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents:
42544
diff
changeset
|
464 |
chop num_ty_args us |>> append mangled_us |
| 43093 | 465 |
val term_ts = map (do_term [] NONE) term_us |
|
42549
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents:
42544
diff
changeset
|
466 |
val T = |
|
43183
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
467 |
(if not (null type_us) andalso |
|
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
468 |
num_type_args thy s' = length type_us then |
|
43191
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
469 |
let val Ts = type_us |> map (typ_from_atp ctxt) in |
|
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
470 |
if new_skolem then |
|
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
471 |
SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) |
|
43200
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents:
43199
diff
changeset
|
472 |
else if textual then |
|
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents:
43199
diff
changeset
|
473 |
try (Sign.const_instance thy) (s', Ts) |
|
43191
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
474 |
else |
|
43200
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents:
43199
diff
changeset
|
475 |
NONE |
|
43191
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
476 |
end |
|
43183
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
477 |
else |
|
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
478 |
NONE) |
|
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
479 |
|> (fn SOME T => T |
|
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
480 |
| NONE => map slack_fastype_of term_ts ---> |
|
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
481 |
(case opt_T of |
|
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
482 |
SOME T => T |
|
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents:
43182
diff
changeset
|
483 |
| NONE => HOLogic.typeT)) |
|
43191
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
484 |
val t = |
|
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
485 |
if new_skolem then |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
486 |
Var ((new_skolem_var_name_from_const s'', var_index), T) |
|
43191
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
487 |
else |
|
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
488 |
Const (unproxify_const s', T) |
|
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
blanchet
parents:
43184
diff
changeset
|
489 |
in list_comb (t, term_ts @ extra_ts) end |
|
42549
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents:
42544
diff
changeset
|
490 |
end |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
491 |
| NONE => (* a free or schematic variable *) |
|
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
492 |
let |
|
45042
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents:
44783
diff
changeset
|
493 |
val term_ts = map (do_term [] NONE) us |
|
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents:
44783
diff
changeset
|
494 |
val ts = term_ts @ extra_ts |
|
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents:
44783
diff
changeset
|
495 |
val T = |
|
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents:
44783
diff
changeset
|
496 |
case opt_T of |
|
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents:
44783
diff
changeset
|
497 |
SOME T => map slack_fastype_of term_ts ---> T |
|
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents:
44783
diff
changeset
|
498 |
| NONE => map slack_fastype_of ts ---> HOLogic.typeT |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
499 |
val t = |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
500 |
case unprefix_and_unascii fixed_var_prefix s of |
|
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
501 |
SOME s => Free (s, T) |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
502 |
| NONE => |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45379
diff
changeset
|
503 |
case unprefix_and_unascii schematic_var_prefix s of |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
504 |
SOME s => Var ((s, var_index), T) |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
505 |
| NONE => |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
506 |
Var ((s |> textual ? repair_variable_name Char.toLower, |
| 43095 | 507 |
var_index), T) |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
508 |
in list_comb (t, ts) end |
| 43093 | 509 |
in do_term [] end |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
510 |
|
|
48132
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents:
48085
diff
changeset
|
511 |
fun term_from_atom ctxt textual sym_tab pos (u as ATerm ((s, _), _)) = |
| 38014 | 512 |
if String.isPrefix class_prefix s then |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
513 |
add_type_constraint pos (type_constraint_from_term ctxt u) |
| 38014 | 514 |
#> pair @{const True}
|
515 |
else |
|
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
516 |
pair (term_from_atp ctxt textual sym_tab (SOME @{typ bool}) u)
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
517 |
|
|
36555
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents:
36551
diff
changeset
|
518 |
val combinator_table = |
| 46904 | 519 |
[(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
|
520 |
(@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
|
|
521 |
(@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
|
|
522 |
(@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
|
|
523 |
(@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
|
|
|
36555
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents:
36551
diff
changeset
|
524 |
|
|
42761
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
525 |
fun uncombine_term thy = |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
526 |
let |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
527 |
fun aux (t1 $ t2) = betapply (pairself aux (t1, t2)) |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
528 |
| aux (Abs (s, T, t')) = Abs (s, T, aux t') |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
529 |
| aux (t as Const (x as (s, _))) = |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
530 |
(case AList.lookup (op =) combinator_table s of |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
531 |
SOME thm => thm |> prop_of |> specialize_type thy x |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
532 |
|> Logic.dest_equals |> snd |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
533 |
| NONE => t) |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
534 |
| aux t = t |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
535 |
in aux end |
|
36555
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents:
36551
diff
changeset
|
536 |
|
| 37991 | 537 |
(* Update schematic type variables with detected sort constraints. It's not |
| 42563 | 538 |
totally clear whether this code is necessary. *) |
| 38014 | 539 |
fun repair_tvar_sorts (t, tvar_tab) = |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
540 |
let |
| 37991 | 541 |
fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) |
542 |
| do_type (TVar (xi, s)) = |
|
543 |
TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) |
|
544 |
| do_type (TFree z) = TFree z |
|
545 |
fun do_term (Const (a, T)) = Const (a, do_type T) |
|
546 |
| do_term (Free (a, T)) = Free (a, do_type T) |
|
547 |
| do_term (Var (xi, T)) = Var (xi, do_type T) |
|
548 |
| do_term (t as Bound _) = t |
|
549 |
| do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) |
|
550 |
| do_term (t1 $ t2) = do_term t1 $ do_term t2 |
|
551 |
in t |> not (Vartab.is_empty tvar_tab) ? do_term end |
|
552 |
||
| 39425 | 553 |
fun quantify_over_var quant_of var_s t = |
554 |
let |
|
555 |
val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s) |
|
556 |
|> map Var |
|
557 |
in fold_rev quant_of vars t end |
|
| 37991 | 558 |
|
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38066
diff
changeset
|
559 |
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they |
|
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38066
diff
changeset
|
560 |
appear in the formula. *) |
| 43184 | 561 |
fun prop_from_atp ctxt textual sym_tab phi = |
| 38014 | 562 |
let |
563 |
fun do_formula pos phi = |
|
| 37991 | 564 |
case phi of |
| 38014 | 565 |
AQuant (_, [], phi) => do_formula pos phi |
| 42526 | 566 |
| AQuant (q, (s, _) :: xs, phi') => |
| 38014 | 567 |
do_formula pos (AQuant (q, xs, phi')) |
| 42526 | 568 |
(* FIXME: TFF *) |
| 39425 | 569 |
#>> quantify_over_var (case q of |
570 |
AForall => forall_of |
|
571 |
| AExists => exists_of) |
|
| 43095 | 572 |
(s |> textual ? repair_variable_name Char.toLower) |
| 38014 | 573 |
| AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not |
| 37991 | 574 |
| AConn (c, [phi1, phi2]) => |
| 38014 | 575 |
do_formula (pos |> c = AImplies ? not) phi1 |
576 |
##>> do_formula pos phi2 |
|
577 |
#>> (case c of |
|
578 |
AAnd => s_conj |
|
579 |
| AOr => s_disj |
|
580 |
| AImplies => s_imp |
|
| 38038 | 581 |
| AIff => s_iff |
| 43163 | 582 |
| ANot => raise Fail "impossible connective") |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
583 |
| AAtom tm => term_from_atom ctxt textual sym_tab pos tm |
| 37991 | 584 |
| _ => raise FORMULA [phi] |
| 38014 | 585 |
in repair_tvar_sorts (do_formula true phi Vartab.empty) end |
| 37991 | 586 |
|
|
43131
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
587 |
fun infer_formula_types ctxt = |
| 39288 | 588 |
Type.constraint HOLogic.boolT |
|
42761
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
589 |
#> Syntax.check_term |
|
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents:
42759
diff
changeset
|
590 |
(Proof_Context.set_mode Proof_Context.mode_schematic ctxt) |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
591 |
|
| 43184 | 592 |
fun uncombined_etc_prop_from_atp ctxt textual sym_tab = |
|
43136
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents:
43135
diff
changeset
|
593 |
let val thy = Proof_Context.theory_of ctxt in |
| 43184 | 594 |
prop_from_atp ctxt textual sym_tab |
|
43176
29a3a1a7794d
only uncombine combinators in textual Isar proofs, not in Metis
blanchet
parents:
43168
diff
changeset
|
595 |
#> textual ? uncombine_term thy #> infer_formula_types ctxt |
|
43136
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents:
43135
diff
changeset
|
596 |
end |
|
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents:
43135
diff
changeset
|
597 |
|
| 43093 | 598 |
(**** Translation of TSTP files to Isar proofs ****) |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
599 |
|
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
600 |
fun unvarify_term (Var ((s, 0), T)) = Free (s, T) |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
601 |
| unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
|
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
602 |
|
| 47774 | 603 |
fun decode_line sym_tab (Definition_Step (name, phi1, phi2)) ctxt = |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
604 |
let |
| 42361 | 605 |
val thy = Proof_Context.theory_of ctxt |
| 43184 | 606 |
val t1 = prop_from_atp ctxt true sym_tab phi1 |
| 36551 | 607 |
val vars = snd (strip_comb t1) |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
608 |
val frees = map unvarify_term vars |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
609 |
val unvarify_args = subst_atomic (vars ~~ frees) |
| 43184 | 610 |
val t2 = prop_from_atp ctxt true sym_tab phi2 |
| 36551 | 611 |
val (t1, t2) = |
612 |
HOLogic.eq_const HOLogic.typeT $ t1 $ t2 |
|
|
43131
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
613 |
|> unvarify_args |> uncombine_term thy |> infer_formula_types ctxt |
|
36555
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents:
36551
diff
changeset
|
614 |
|> HOLogic.dest_eq |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
615 |
in |
| 47774 | 616 |
(Definition_Step (name, t1, t2), |
| 44121 | 617 |
fold Variable.declare_term (maps Misc_Legacy.term_frees [t1, t2]) ctxt) |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
618 |
end |
| 47774 | 619 |
| decode_line sym_tab (Inference_Step (name, u, rule, deps)) ctxt = |
| 43184 | 620 |
let val t = u |> uncombined_etc_prop_from_atp ctxt true sym_tab in |
| 47774 | 621 |
(Inference_Step (name, t, rule, deps), |
| 44121 | 622 |
fold Variable.declare_term (Misc_Legacy.term_frees t) ctxt) |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
623 |
end |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
624 |
fun decode_lines ctxt sym_tab lines = |
|
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
625 |
fst (fold_map (decode_line sym_tab) lines ctxt) |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
626 |
|
| 47774 | 627 |
fun is_same_inference _ (Definition_Step _) = false |
628 |
| is_same_inference t (Inference_Step (_, t', _, _)) = t aconv t' |
|
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
629 |
|
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
630 |
(* No "real" literals means only type information (tfree_tcs, clsrel, or |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
631 |
clsarity). *) |
| 47918 | 632 |
fun is_only_type_information t = t aconv @{term True}
|
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
633 |
|
| 39373 | 634 |
fun replace_one_dependency (old, new) dep = |
|
42968
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents:
42966
diff
changeset
|
635 |
if is_same_atp_step dep old then new else [dep] |
| 47774 | 636 |
fun replace_dependencies_in_line _ (line as Definition_Step _) = line |
637 |
| replace_dependencies_in_line p (Inference_Step (name, t, rule, deps)) = |
|
638 |
Inference_Step (name, t, rule, |
|
639 |
fold (union (op =) o replace_one_dependency p) deps []) |
|
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
640 |
|
|
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40114
diff
changeset
|
641 |
(* Discard facts; consolidate adjacent lines that prove the same formula, since |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38066
diff
changeset
|
642 |
they differ only in type information.*) |
| 47774 | 643 |
fun add_line _ (line as Definition_Step _) lines = line :: lines |
644 |
| add_line fact_names (Inference_Step (name as (_, ss), t, rule, [])) lines = |
|
|
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40114
diff
changeset
|
645 |
(* No dependencies: fact, conjecture, or (for Vampire) internal facts or |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38066
diff
changeset
|
646 |
definitions. *) |
| 45552 | 647 |
if is_fact fact_names ss then |
|
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40114
diff
changeset
|
648 |
(* Facts are not proof lines. *) |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
649 |
if is_only_type_information t then |
| 39373 | 650 |
map (replace_dependencies_in_line (name, [])) lines |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
651 |
(* Is there a repetition? If so, replace later line by earlier one. *) |
| 38035 | 652 |
else case take_prefix (not o is_same_inference t) lines of |
| 39373 | 653 |
(_, []) => lines (* no repetition of proof line *) |
| 47774 | 654 |
| (pre, Inference_Step (name', _, _, _) :: post) => |
| 39373 | 655 |
pre @ map (replace_dependencies_in_line (name', [name])) post |
| 40069 | 656 |
| _ => raise Fail "unexpected inference" |
| 45552 | 657 |
else if is_conjecture ss then |
| 47918 | 658 |
Inference_Step (name, t, rule, []) :: lines |
| 36551 | 659 |
else |
| 39373 | 660 |
map (replace_dependencies_in_line (name, [])) lines |
| 47774 | 661 |
| add_line _ (Inference_Step (name, t, rule, deps)) lines = |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
662 |
(* Type information will be deleted later; skip repetition test. *) |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
663 |
if is_only_type_information t then |
| 47774 | 664 |
Inference_Step (name, t, rule, deps) :: lines |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
665 |
(* Is there a repetition? If so, replace later line by earlier one. *) |
| 38035 | 666 |
else case take_prefix (not o is_same_inference t) lines of |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
667 |
(* FIXME: Doesn't this code risk conflating proofs involving different |
| 38035 | 668 |
types? *) |
| 47774 | 669 |
(_, []) => Inference_Step (name, t, rule, deps) :: lines |
670 |
| (pre, Inference_Step (name', t', rule, _) :: post) => |
|
671 |
Inference_Step (name, t', rule, deps) :: |
|
| 39373 | 672 |
pre @ map (replace_dependencies_in_line (name', [name])) post |
| 40069 | 673 |
| _ => raise Fail "unexpected inference" |
|
22044
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
674 |
|
|
47927
c35238d19bb9
repair the Waldmeister endgame only for Waldmeister proofs
blanchet
parents:
47921
diff
changeset
|
675 |
val waldmeister_conjecture_num = "1.0.0.0" |
|
c35238d19bb9
repair the Waldmeister endgame only for Waldmeister proofs
blanchet
parents:
47921
diff
changeset
|
676 |
|
| 47918 | 677 |
val repair_waldmeister_endgame = |
678 |
let |
|
679 |
fun do_tail (Inference_Step (name, t, rule, deps)) = |
|
680 |
Inference_Step (name, s_not t, rule, deps) |
|
681 |
| do_tail line = line |
|
682 |
fun do_body [] = [] |
|
|
47927
c35238d19bb9
repair the Waldmeister endgame only for Waldmeister proofs
blanchet
parents:
47921
diff
changeset
|
683 |
| do_body ((line as Inference_Step ((num, _), _, _, _)) :: lines) = |
|
c35238d19bb9
repair the Waldmeister endgame only for Waldmeister proofs
blanchet
parents:
47921
diff
changeset
|
684 |
if num = waldmeister_conjecture_num then map do_tail (line :: lines) |
| 47918 | 685 |
else line :: do_body lines |
686 |
| do_body (line :: lines) = line :: do_body lines |
|
687 |
in do_body end |
|
688 |
||
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
689 |
(* Recursively delete empty lines (type information) from the proof. *) |
| 47774 | 690 |
fun add_nontrivial_line (line as Inference_Step (name, t, _, [])) lines = |
| 39373 | 691 |
if is_only_type_information t then delete_dependency name lines |
| 45209 | 692 |
else line :: lines |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
693 |
| add_nontrivial_line line lines = line :: lines |
| 39373 | 694 |
and delete_dependency name lines = |
695 |
fold_rev add_nontrivial_line |
|
696 |
(map (replace_dependencies_in_line (name, [])) lines) [] |
|
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
697 |
|
| 37323 | 698 |
(* ATPs sometimes reuse free variable names in the strangest ways. Removing |
699 |
offending lines often does the trick. *) |
|
|
36560
45c1870f234f
fixed definition of "bad frees" so that it actually works
blanchet
parents:
36559
diff
changeset
|
700 |
fun is_bad_free frees (Free x) = not (member (op =) frees x) |
|
45c1870f234f
fixed definition of "bad frees" so that it actually works
blanchet
parents:
36559
diff
changeset
|
701 |
| is_bad_free _ _ = false |
|
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
702 |
|
| 47774 | 703 |
fun add_desired_line _ _ _ (line as Definition_Step (name, _, _)) (j, lines) = |
| 39373 | 704 |
(j, line :: map (replace_dependencies_in_line (name, [])) lines) |
| 45551 | 705 |
| add_desired_line isar_shrink_factor fact_names frees |
| 47774 | 706 |
(Inference_Step (name as (_, ss), t, rule, deps)) (j, lines) = |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
707 |
(j + 1, |
| 45552 | 708 |
if is_fact fact_names ss orelse |
709 |
is_conjecture ss orelse |
|
| 39373 | 710 |
(* the last line must be kept *) |
711 |
j = 0 orelse |
|
|
36570
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents:
36567
diff
changeset
|
712 |
(not (is_only_type_information t) andalso |
|
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents:
36567
diff
changeset
|
713 |
null (Term.add_tvars t []) andalso |
|
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents:
36567
diff
changeset
|
714 |
not (exists_subterm (is_bad_free frees) t) andalso |
| 39373 | 715 |
length deps >= 2 andalso j mod isar_shrink_factor = 0 andalso |
716 |
(* kill next to last line, which usually results in a trivial step *) |
|
717 |
j <> 1) then |
|
| 47774 | 718 |
Inference_Step (name, t, rule, deps) :: lines (* keep line *) |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
719 |
else |
| 39373 | 720 |
map (replace_dependencies_in_line (name, deps)) lines) (* drop line *) |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
721 |
|
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
722 |
(** Isar proof construction and manipulation **) |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
723 |
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
724 |
type label = string * int |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
725 |
type facts = label list * string list |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
726 |
|
| 39452 | 727 |
datatype isar_qualifier = Show | Then | Moreover | Ultimately |
|
36291
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
blanchet
parents:
36288
diff
changeset
|
728 |
|
| 39452 | 729 |
datatype isar_step = |
|
36478
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36477
diff
changeset
|
730 |
Fix of (string * typ) list | |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
731 |
Let of term * term | |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
732 |
Assume of label * term | |
| 45882 | 733 |
Prove of isar_qualifier list * label * term * byline |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
734 |
and byline = |
| 45882 | 735 |
By_Metis of facts | |
736 |
Case_Split of isar_step list list * facts |
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
737 |
|
| 45552 | 738 |
fun add_fact_from_dependency fact_names (name as (_, ss)) = |
739 |
if is_fact fact_names ss then |
|
740 |
apsnd (union (op =) (map fst (resolve_fact fact_names ss))) |
|
|
36475
05209b869a6b
new Isar proof construction code: stringfy axiom names correctly
blanchet
parents:
36474
diff
changeset
|
741 |
else |
| 45551 | 742 |
apfst (insert (op =) (raw_label_for_name name)) |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
743 |
|
| 39454 | 744 |
fun repair_name "$true" = "c_True" |
745 |
| repair_name "$false" = "c_False" |
|
|
43000
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
blanchet
parents:
42998
diff
changeset
|
746 |
| repair_name "$$e" = tptp_equal (* seen in Vampire proofs *) |
| 39454 | 747 |
| repair_name s = |
|
43000
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
blanchet
parents:
42998
diff
changeset
|
748 |
if is_tptp_equal s orelse |
|
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
blanchet
parents:
42998
diff
changeset
|
749 |
(* seen in Vampire proofs *) |
|
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
blanchet
parents:
42998
diff
changeset
|
750 |
(String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then |
|
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
blanchet
parents:
42998
diff
changeset
|
751 |
tptp_equal |
| 39454 | 752 |
else |
753 |
s |
|
754 |
||
| 45883 | 755 |
(* FIXME: Still needed? Try with SPASS proofs perhaps. *) |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
756 |
val kill_duplicate_assumptions_in_proof = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
757 |
let |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
758 |
fun relabel_facts subst = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
759 |
apfst (map (fn l => AList.lookup (op =) subst l |> the_default l)) |
|
36491
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents:
36488
diff
changeset
|
760 |
fun do_step (step as Assume (l, t)) (proof, subst, assums) = |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
761 |
(case AList.lookup (op aconv) assums t of |
|
36967
3c804030474b
fix bug in Isar proof reconstruction step relabeling + don't try to infer the sorts of TVars, since this often fails miserably
blanchet
parents:
36924
diff
changeset
|
762 |
SOME l' => (proof, (l, l') :: subst, assums) |
|
36491
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents:
36488
diff
changeset
|
763 |
| NONE => (step :: proof, subst, (t, l) :: assums)) |
| 45882 | 764 |
| do_step (Prove (qs, l, t, by)) (proof, subst, assums) = |
765 |
(Prove (qs, l, t, |
|
766 |
case by of |
|
767 |
By_Metis facts => By_Metis (relabel_facts subst facts) |
|
768 |
| Case_Split (proofs, facts) => |
|
769 |
Case_Split (map do_proof proofs, |
|
770 |
relabel_facts subst facts)) :: |
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
771 |
proof, subst, assums) |
|
36491
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents:
36488
diff
changeset
|
772 |
| do_step step (proof, subst, assums) = (step :: proof, subst, assums) |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
773 |
and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
774 |
in do_proof end |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
775 |
|
| 45883 | 776 |
fun used_labels_of_step (Prove (_, _, _, by)) = |
777 |
(case by of |
|
778 |
By_Metis (ls, _) => ls |
|
779 |
| Case_Split (proofs, (ls, _)) => |
|
780 |
fold (union (op =) o used_labels_of) proofs ls) |
|
781 |
| used_labels_of_step _ = [] |
|
782 |
and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof [] |
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
783 |
|
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
784 |
fun kill_useless_labels_in_proof proof = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
785 |
let |
|
36556
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36555
diff
changeset
|
786 |
val used_ls = used_labels_of proof |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
787 |
fun do_label l = if member (op =) used_ls l then l else no_label |
|
36556
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36555
diff
changeset
|
788 |
fun do_step (Assume (l, t)) = Assume (do_label l, t) |
| 45882 | 789 |
| do_step (Prove (qs, l, t, by)) = |
790 |
Prove (qs, do_label l, t, |
|
791 |
case by of |
|
792 |
Case_Split (proofs, facts) => |
|
793 |
Case_Split (map (map do_step) proofs, facts) |
|
794 |
| _ => by) |
|
|
36556
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36555
diff
changeset
|
795 |
| do_step step = step |
|
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36555
diff
changeset
|
796 |
in map do_step proof end |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
797 |
|
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
798 |
fun prefix_for_depth n = replicate_string (n + 1) |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
799 |
|
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
800 |
val relabel_proof = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
801 |
let |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
802 |
fun aux _ _ _ [] = [] |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
803 |
| aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
804 |
if l = no_label then |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
805 |
Assume (l, t) :: aux subst depth (next_assum, next_fact) proof |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
806 |
else |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
807 |
let val l' = (prefix_for_depth depth assum_prefix, next_assum) in |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
808 |
Assume (l', t) :: |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
809 |
aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
810 |
end |
| 45882 | 811 |
| aux subst depth (next_assum, next_fact) |
812 |
(Prove (qs, l, t, by) :: proof) = |
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
813 |
let |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
814 |
val (l', subst, next_fact) = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
815 |
if l = no_label then |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
816 |
(l, subst, next_fact) |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
817 |
else |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
818 |
let |
|
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41742
diff
changeset
|
819 |
val l' = (prefix_for_depth depth have_prefix, next_fact) |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
820 |
in (l', (l, l') :: subst, next_fact + 1) end |
|
36570
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents:
36567
diff
changeset
|
821 |
val relabel_facts = |
|
39370
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents:
39368
diff
changeset
|
822 |
apfst (maps (the_list o AList.lookup (op =) subst)) |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
823 |
val by = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
824 |
case by of |
| 45882 | 825 |
By_Metis facts => By_Metis (relabel_facts facts) |
826 |
| Case_Split (proofs, facts) => |
|
827 |
Case_Split (map (aux subst (depth + 1) (1, 1)) proofs, |
|
828 |
relabel_facts facts) |
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
829 |
in |
| 45882 | 830 |
Prove (qs, l', t, by) :: aux subst depth (next_assum, next_fact) proof |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
831 |
end |
|
36491
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents:
36488
diff
changeset
|
832 |
| aux subst depth nextp (step :: proof) = |
|
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents:
36488
diff
changeset
|
833 |
step :: aux subst depth nextp proof |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
834 |
in aux [] 0 (1, 1) end |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
835 |
|
|
49740
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
836 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
837 |
(** Type annotations **) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
838 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
839 |
fun post_traverse_term_type' f _ (t as Const (_, T)) s = f t T s |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
840 |
| post_traverse_term_type' f _ (t as Free (_, T)) s = f t T s |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
841 |
| post_traverse_term_type' f _ (t as Var (_, T)) s = f t T s |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
842 |
| post_traverse_term_type' f env (t as Bound i) s = f t (nth env i) s |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
843 |
| post_traverse_term_type' f env (Abs (x, T1, b)) s = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
844 |
let |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
845 |
val ((b', s'), T2) = post_traverse_term_type' f (T1 :: env) b s |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
846 |
in f (Abs (x, T1, b')) (T1 --> T2) s' end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
847 |
| post_traverse_term_type' f env (u $ v) s = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
848 |
let |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
849 |
val ((u', s'), Type (_, [_, T])) = post_traverse_term_type' f env u s |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
850 |
val ((v', s''), _) = post_traverse_term_type' f env v s' |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
851 |
in f (u' $ v') T s'' end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
852 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
853 |
fun post_traverse_term_type f s t = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
854 |
post_traverse_term_type' (fn t => fn T => fn s => (f t T s, T)) [] t s |> fst |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
855 |
fun post_fold_term_type f s t = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
856 |
post_traverse_term_type (fn t => fn T => fn s => (t, f t T s)) s t |> snd |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
857 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
858 |
(* Data structures, orders *) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
859 |
val cost_ord = prod_ord int_ord (prod_ord int_ord int_ord) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
860 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
861 |
structure Var_Set_Tab = Table( |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
862 |
type key = indexname list |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
863 |
val ord = list_ord Term_Ord.fast_indexname_ord) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
864 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
865 |
(* (1) Generalize Types *) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
866 |
fun generalize_types ctxt t = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
867 |
t |> map_types (fn _ => dummyT) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
868 |
|> Syntax.check_term |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
869 |
(Proof_Context.set_mode Proof_Context.mode_pattern ctxt) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
870 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
871 |
(* (2) Typing-spot Table *) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
872 |
local |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
873 |
fun key_of_atype (TVar (idxn, _)) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
874 |
Ord_List.insert Term_Ord.fast_indexname_ord idxn |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
875 |
| key_of_atype _ = I |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
876 |
fun key_of_type T = fold_atyps key_of_atype T [] |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
877 |
fun update_tab t T (tab, pos) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
878 |
(case key_of_type T of |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
879 |
[] => tab |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
880 |
| key => |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
881 |
let val cost = (size_of_typ T, (size_of_term t, pos)) in |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
882 |
case Var_Set_Tab.lookup tab key of |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
883 |
NONE => Var_Set_Tab.update_new (key, cost) tab |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
884 |
| SOME old_cost => |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
885 |
(case cost_ord (cost, old_cost) of |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
886 |
LESS => Var_Set_Tab.update (key, cost) tab |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
887 |
| _ => tab) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
888 |
end, |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
889 |
pos + 1) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
890 |
in |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
891 |
val typing_spot_table = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
892 |
post_fold_term_type update_tab (Var_Set_Tab.empty, 0) #> fst |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
893 |
end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
894 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
895 |
(* (3) Reverse-Greedy *) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
896 |
fun reverse_greedy typing_spot_tab = |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
897 |
let |
|
49740
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
898 |
fun update_count z = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
899 |
fold (fn tvar => fn tab => |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
900 |
let val c = Vartab.lookup tab tvar |> the_default 0 in |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
901 |
Vartab.update (tvar, c + z) tab |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
902 |
end) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
903 |
fun superfluous tcount = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
904 |
forall (fn tvar => the (Vartab.lookup tcount tvar) > 1) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
905 |
fun drop_superfluous (tvars, (_, (_, spot))) (spots, tcount) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
906 |
if superfluous tcount tvars then (spots, update_count ~1 tvars tcount) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
907 |
else (spot :: spots, tcount) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
908 |
val (typing_spots, tvar_count_tab) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
909 |
Var_Set_Tab.fold |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
910 |
(fn kv as (k, _) => apfst (cons kv) #> apsnd (update_count 1 k)) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
911 |
typing_spot_tab ([], Vartab.empty) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
912 |
|>> sort_distinct (rev_order o cost_ord o pairself snd) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
913 |
in fold drop_superfluous typing_spots ([], tvar_count_tab) |> fst end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
914 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
915 |
(* (4) Introduce Annotations *) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
916 |
fun introduce_annotations thy spots t t' = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
917 |
let |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
918 |
val get_types = post_fold_term_type (K cons) [] |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
919 |
fun match_types tp = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
920 |
fold (Sign.typ_match thy) (op ~~ (pairself get_types tp)) Vartab.empty |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
921 |
fun unica' b x [] = if b then [x] else [] |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
922 |
| unica' b x (y :: ys) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
923 |
if x = y then unica' false x ys |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
924 |
else unica' true y ys |> b ? cons x |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
925 |
fun unica ord xs = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
926 |
case sort ord xs of x :: ys => unica' true x ys | [] => [] |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
927 |
val add_all_tfree_namesT = fold_atyps (fn TFree (x, _) => cons x | _ => I) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
928 |
fun erase_unica_tfrees env = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
929 |
let |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
930 |
val unica = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
931 |
Vartab.fold (add_all_tfree_namesT o snd o snd) env [] |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
932 |
|> unica fast_string_ord |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
933 |
val erase_unica = map_atyps |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
934 |
(fn T as TFree (s, _) => |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
935 |
if Ord_List.member fast_string_ord unica s then dummyT else T |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
936 |
| T => T) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
937 |
in Vartab.map (K (apsnd erase_unica)) env end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
938 |
val env = match_types (t', t) |> erase_unica_tfrees |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
939 |
fun get_annot env (TFree _) = (false, (env, dummyT)) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
940 |
| get_annot env (T as TVar (v, S)) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
941 |
let val T' = Envir.subst_type env T in |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
942 |
if T' = dummyT then (false, (env, dummyT)) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
943 |
else (true, (Vartab.update (v, (S, dummyT)) env, T')) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
944 |
end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
945 |
| get_annot env (Type (S, Ts)) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
946 |
(case fold_rev (fn T => fn (b, (env, Ts)) => |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
947 |
let |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
948 |
val (b', (env', T)) = get_annot env T |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
949 |
in (b orelse b', (env', T :: Ts)) end) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
950 |
Ts (false, (env, [])) of |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
951 |
(true, (env', Ts)) => (true, (env', Type (S, Ts))) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
952 |
| (false, (env', _)) => (false, (env', dummyT))) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
953 |
fun post1 _ T (env, cp, ps as p :: ps', annots) = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
954 |
if p <> cp then |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
955 |
(env, cp + 1, ps, annots) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
956 |
else |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
957 |
let val (_, (env', T')) = get_annot env T in |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
958 |
(env', cp + 1, ps', (p, T') :: annots) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
959 |
end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
960 |
| post1 _ _ accum = accum |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
961 |
val (_, _, _, annots) = post_fold_term_type post1 (env, 0, spots, []) t' |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
962 |
fun post2 t _ (cp, annots as (p, T) :: annots') = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
963 |
if p <> cp then (t, (cp + 1, annots)) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
964 |
else (Type.constraint T t, (cp + 1, annots')) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
965 |
| post2 t _ x = (t, x) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
966 |
in post_traverse_term_type post2 (0, rev annots) t |> fst end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
967 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
968 |
(* (5) Annotate *) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
969 |
fun annotate_types ctxt t = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
970 |
let |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
971 |
val thy = Proof_Context.theory_of ctxt |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
972 |
val t' = generalize_types ctxt t |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
973 |
val typing_spots = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
974 |
t' |> typing_spot_table |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
975 |
|> reverse_greedy |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
976 |
|> sort int_ord |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
977 |
in introduce_annotations thy typing_spots t t' end |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
978 |
|
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
979 |
fun string_for_proof ctxt type_enc lam_trans i n = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
980 |
let |
|
37319
42268dc7d6c4
show types in Isar proofs, but not for free variables;
blanchet
parents:
37172
diff
changeset
|
981 |
fun fix_print_mode f x = |
|
39134
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
wenzelm
parents:
39115
diff
changeset
|
982 |
Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) |
|
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
wenzelm
parents:
39115
diff
changeset
|
983 |
(print_mode_value ())) f x |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
984 |
fun do_indent ind = replicate_string (ind * indent_size) " " |
|
36478
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36477
diff
changeset
|
985 |
fun do_free (s, T) = |
|
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36477
diff
changeset
|
986 |
maybe_quote s ^ " :: " ^ |
|
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36477
diff
changeset
|
987 |
maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T) |
|
36570
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents:
36567
diff
changeset
|
988 |
fun do_label l = if l = no_label then "" else string_for_label l ^ ": " |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
989 |
fun do_have qs = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
990 |
(if member (op =) qs Moreover then "moreover " else "") ^ |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
991 |
(if member (op =) qs Ultimately then "ultimately " else "") ^ |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
992 |
(if member (op =) qs Then then |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
993 |
if member (op =) qs Show then "thus" else "hence" |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
994 |
else |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
995 |
if member (op =) qs Show then "show" else "have") |
|
49740
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
996 |
val do_term = |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
997 |
maybe_quote o fix_print_mode (Syntax.string_of_term ctxt) |
|
6f02b893dd87
added type annotations to generated Isar proofs -- code by Steffen Smolka
blanchet
parents:
48799
diff
changeset
|
998 |
o annotate_types ctxt |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
999 |
val reconstr = Metis (type_enc, lam_trans) |
|
36570
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents:
36567
diff
changeset
|
1000 |
fun do_facts (ls, ss) = |
| 48799 | 1001 |
reconstructor_command reconstr 1 1 [] 0 |
| 43033 | 1002 |
(ls |> sort_distinct (prod_ord string_ord int_ord), |
1003 |
ss |> sort_distinct string_ord) |
|
|
36478
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36477
diff
changeset
|
1004 |
and do_step ind (Fix xs) = |
|
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36477
diff
changeset
|
1005 |
do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n" |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
1006 |
| do_step ind (Let (t1, t2)) = |
|
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
1007 |
do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n" |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1008 |
| do_step ind (Assume (l, t)) = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1009 |
do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n" |
| 45882 | 1010 |
| do_step ind (Prove (qs, l, t, By_Metis facts)) = |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1011 |
do_indent ind ^ do_have qs ^ " " ^ |
| 36479 | 1012 |
do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n" |
| 45882 | 1013 |
| do_step ind (Prove (qs, l, t, Case_Split (proofs, facts))) = |
1014 |
implode (map (prefix (do_indent ind ^ "moreover\n") o do_block ind) |
|
1015 |
proofs) ^ |
|
| 36479 | 1016 |
do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^ |
|
36478
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36477
diff
changeset
|
1017 |
do_facts facts ^ "\n" |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1018 |
and do_steps prefix suffix ind steps = |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1019 |
let val s = implode (map (do_step ind) steps) in |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1020 |
replicate_string (ind * indent_size - size prefix) " " ^ prefix ^ |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1021 |
String.extract (s, ind * indent_size, |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1022 |
SOME (size s - ind * indent_size - 1)) ^ |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1023 |
suffix ^ "\n" |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1024 |
end |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1025 |
and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
|
|
36564
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
blanchet
parents:
36563
diff
changeset
|
1026 |
(* One-step proofs are pointless; better use the Metis one-liner |
|
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
blanchet
parents:
36563
diff
changeset
|
1027 |
directly. *) |
| 45882 | 1028 |
and do_proof [Prove (_, _, _, By_Metis _)] = "" |
|
36564
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
blanchet
parents:
36563
diff
changeset
|
1029 |
| do_proof proof = |
|
36480
1e01a7162435
polish Isar proofs: don't mention facts twice, and don't show one-liner "structured" proofs
blanchet
parents:
36479
diff
changeset
|
1030 |
(if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ |
| 39452 | 1031 |
do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^ |
1032 |
(if n <> 1 then "next" else "qed") |
|
|
36488
32c92af68ec9
remove Sledgehammer's "sorts" option to annotate variables with sorts in proof;
blanchet
parents:
36486
diff
changeset
|
1033 |
in do_proof end |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1034 |
|
|
43062
2834548a7a48
nicer failure message when one-line proof reconstruction fails
blanchet
parents:
43051
diff
changeset
|
1035 |
fun isar_proof_text ctxt isar_proof_requested |
| 45552 | 1036 |
(debug, isar_shrink_factor, pool, fact_names, sym_tab, atp_proof, goal) |
| 43037 | 1037 |
(one_line_params as (_, _, _, _, subgoal, subgoal_count)) = |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1038 |
let |
|
43062
2834548a7a48
nicer failure message when one-line proof reconstruction fails
blanchet
parents:
43051
diff
changeset
|
1039 |
val isar_shrink_factor = |
|
2834548a7a48
nicer failure message when one-line proof reconstruction fails
blanchet
parents:
43051
diff
changeset
|
1040 |
(if isar_proof_requested then 1 else 2) * isar_shrink_factor |
| 43037 | 1041 |
val (params, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
1042 |
val frees = fold Term.add_frees (concl_t :: hyp_ts) [] |
| 48799 | 1043 |
val one_line_proof = one_line_proof_text 0 one_line_params |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
1044 |
val type_enc = |
| 45590 | 1045 |
if is_typed_helper_used_in_atp_proof atp_proof then full_typesN |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
1046 |
else partial_typesN |
|
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
1047 |
val lam_trans = lam_trans_from_atp_proof atp_proof metis_default_lam_trans |
| 45882 | 1048 |
|
| 45883 | 1049 |
fun isar_proof_of () = |
1050 |
let |
|
1051 |
val atp_proof = |
|
1052 |
atp_proof |
|
1053 |
|> clean_up_atp_proof_dependencies |
|
1054 |
|> nasty_atp_proof pool |
|
1055 |
|> map_term_names_in_atp_proof repair_name |
|
1056 |
|> decode_lines ctxt sym_tab |
|
1057 |
|> rpair [] |-> fold_rev (add_line fact_names) |
|
| 47918 | 1058 |
|> repair_waldmeister_endgame |
| 45883 | 1059 |
|> rpair [] |-> fold_rev add_nontrivial_line |
1060 |
|> rpair (0, []) |
|
1061 |
|-> fold_rev (add_desired_line isar_shrink_factor fact_names frees) |
|
1062 |
|> snd |
|
1063 |
val conj_name = conjecture_prefix ^ string_of_int (length hyp_ts) |
|
1064 |
val conjs = |
|
1065 |
atp_proof |
|
| 47774 | 1066 |
|> map_filter (fn Inference_Step (name as (_, ss), _, _, []) => |
| 45883 | 1067 |
if member (op =) ss conj_name then SOME name else NONE |
1068 |
| _ => NONE) |
|
| 47774 | 1069 |
fun dep_of_step (Definition_Step _) = NONE |
1070 |
| dep_of_step (Inference_Step (name, _, _, from)) = SOME (from, name) |
|
| 45883 | 1071 |
val ref_graph = atp_proof |> map_filter dep_of_step |> make_ref_graph |
1072 |
val axioms = axioms_of_ref_graph ref_graph conjs |
|
1073 |
val tainted = tainted_atoms_of_ref_graph ref_graph conjs |
|
1074 |
val props = |
|
1075 |
Symtab.empty |
|
| 47774 | 1076 |
|> fold (fn Definition_Step _ => I (* FIXME *) |
1077 |
| Inference_Step ((s, _), t, _, _) => |
|
| 45883 | 1078 |
Symtab.update_new (s, |
|
47921
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
blanchet
parents:
47920
diff
changeset
|
1079 |
t |> fold forall_of (map Var (Term.add_vars t [])) |
| 47918 | 1080 |
|> member (op = o apsnd fst) tainted s ? s_not)) |
| 45883 | 1081 |
atp_proof |
1082 |
fun prop_of_clause c = |
|
1083 |
fold (curry s_disj) (map_filter (Symtab.lookup props o fst) c) |
|
1084 |
@{term False}
|
|
|
47920
a5c2386518e2
ensure consistent naming of Waldmeister proof steps, so that they are not cleaned away by "clean_up_atp_proof_dependencies"
blanchet
parents:
47918
diff
changeset
|
1085 |
fun label_of_clause [name] = raw_label_for_name name |
|
a5c2386518e2
ensure consistent naming of Waldmeister proof steps, so that they are not cleaned away by "clean_up_atp_proof_dependencies"
blanchet
parents:
47918
diff
changeset
|
1086 |
| label_of_clause c = (space_implode "___" (map fst c), 0) |
| 45883 | 1087 |
fun maybe_show outer c = |
1088 |
(outer andalso length c = 1 andalso subset (op =) (c, conjs)) |
|
1089 |
? cons Show |
|
1090 |
fun do_have outer qs (gamma, c) = |
|
1091 |
Prove (maybe_show outer c qs, label_of_clause c, prop_of_clause c, |
|
1092 |
By_Metis (fold (add_fact_from_dependency fact_names |
|
1093 |
o the_single) gamma ([], []))) |
|
1094 |
fun do_inf outer (Have z) = do_have outer [] z |
|
1095 |
| do_inf outer (Hence z) = do_have outer [Then] z |
|
1096 |
| do_inf outer (Cases cases) = |
|
1097 |
let val c = succedent_of_cases cases in |
|
1098 |
Prove (maybe_show outer c [Ultimately], label_of_clause c, |
|
1099 |
prop_of_clause c, |
|
1100 |
Case_Split (map (do_case false) cases, ([], []))) |
|
1101 |
end |
|
1102 |
and do_case outer (c, infs) = |
|
1103 |
Assume (label_of_clause c, prop_of_clause c) :: |
|
1104 |
map (do_inf outer) infs |
|
1105 |
val isar_proof = |
|
1106 |
(if null params then [] else [Fix params]) @ |
|
1107 |
(ref_graph |
|
1108 |
|> redirect_graph axioms tainted |
|
1109 |
|> chain_direct_proof |
|
1110 |
|> map (do_inf true) |
|
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1111 |
|> kill_duplicate_assumptions_in_proof |
|
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1112 |
|> kill_useless_labels_in_proof |
| 45883 | 1113 |
|> relabel_proof) |
1114 |
|> string_for_proof ctxt type_enc lam_trans subgoal subgoal_count |
|
1115 |
in |
|
1116 |
case isar_proof of |
|
1117 |
"" => |
|
1118 |
if isar_proof_requested then |
|
1119 |
"\nNo structured proof available (proof too short)." |
|
1120 |
else |
|
1121 |
"" |
|
1122 |
| _ => |
|
1123 |
"\n\n" ^ (if isar_proof_requested then "Structured proof" |
|
1124 |
else "Perhaps this will work") ^ |
|
1125 |
":\n" ^ Markup.markup Isabelle_Markup.sendback isar_proof |
|
1126 |
end |
|
|
35868
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
1127 |
val isar_proof = |
|
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
1128 |
if debug then |
| 45883 | 1129 |
isar_proof_of () |
1130 |
else case try isar_proof_of () of |
|
| 45882 | 1131 |
SOME s => s |
1132 |
| NONE => if isar_proof_requested then |
|
1133 |
"\nWarning: The Isar proof construction failed." |
|
1134 |
else |
|
1135 |
"" |
|
| 43033 | 1136 |
in one_line_proof ^ isar_proof end |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
1137 |
|
| 48799 | 1138 |
fun proof_text ctxt isar_proof isar_params num_chained |
| 43033 | 1139 |
(one_line_params as (preplay, _, _, _, _, _)) = |
| 43166 | 1140 |
(if case preplay of Failed_to_Play _ => true | _ => isar_proof then |
|
43062
2834548a7a48
nicer failure message when one-line proof reconstruction fails
blanchet
parents:
43051
diff
changeset
|
1141 |
isar_proof_text ctxt isar_proof isar_params |
| 43033 | 1142 |
else |
| 48799 | 1143 |
one_line_proof_text num_chained) one_line_params |
|
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
1144 |
|
| 31038 | 1145 |
end; |