| author | desharna | 
| Mon, 25 Mar 2024 19:27:32 +0100 | |
| changeset 79996 | 4f803ae64781 | 
| parent 76183 | 8089593a364a | 
| permissions | -rw-r--r-- | 
| 76183 | 1  | 
(* Title: HOL/Tools/SMT/lethe_proof_parse.ML  | 
| 
75299
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
2  | 
Author: Mathias Fleury, TU Muenchen  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
3  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
5  | 
VeriT proof parsing.  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
8  | 
signature LETHE_PROOF_PARSE =  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
10  | 
  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
 | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
11  | 
val parse_proof: SMT_Translate.replay_data ->  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
12  | 
((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
13  | 
SMT_Solver.parsed_proof  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
14  | 
end;  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
15  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
16  | 
structure Lethe_Proof_Parse: LETHE_PROOF_PARSE =  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
17  | 
struct  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
19  | 
open ATP_Util  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
20  | 
open ATP_Problem  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
21  | 
open ATP_Proof  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
22  | 
open ATP_Proof_Reconstruct  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
23  | 
open Lethe_Isar  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
24  | 
open Lethe_Proof  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
26  | 
fun parse_proof  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
27  | 
    ({context = ctxt, typs, terms, ll_defs, rewrite_rules, assms} : SMT_Translate.replay_data)
 | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
28  | 
xfacts prems concl output =  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
29  | 
let  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
30  | 
val num_ll_defs = length ll_defs  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
31  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
32  | 
val id_of_index = Integer.add num_ll_defs  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
33  | 
val index_of_id = Integer.add (~ num_ll_defs)  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
34  | 
|
| 75367 | 35  | 
fun step_of_assume j ((_, role), th) =  | 
36  | 
Lethe_Proof.Lethe_Step  | 
|
37  | 
        {id = SMTLIB_Interface.assert_name_of_role_and_index role (id_of_index j),
 | 
|
| 
75299
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
38  | 
rule = input_rule, prems = [], proof_ctxt = [], concl = Thm.prop_of th, fixes = []}  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
40  | 
val (actual_steps, _) = Lethe_Proof.parse typs terms output ctxt  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
41  | 
val used_assert_ids =  | 
| 75367 | 42  | 
actual_steps  | 
43  | 
        |> map_filter (fn (Lethe_Step { id, ...}) =>
 | 
|
44  | 
try (snd o SMTLIB_Interface.role_and_index_of_assert_name) id)  | 
|
| 
75299
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
45  | 
val used_assm_js =  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
46  | 
map_filter (fn id => let val i = index_of_id id in if i >= 0 then SOME i else NONE end)  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
47  | 
used_assert_ids  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
48  | 
val used_assms = map (nth assms) used_assm_js  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
49  | 
val assm_steps = map2 step_of_assume used_assm_js used_assms  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
50  | 
val steps = assm_steps @ actual_steps  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
52  | 
val conjecture_i = 0  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
53  | 
val prems_i = conjecture_i + 1  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
54  | 
val num_prems = length prems  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
55  | 
val facts_i = prems_i + num_prems  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
56  | 
val num_facts = length xfacts  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
57  | 
val helpers_i = facts_i + num_facts  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
59  | 
val conjecture_id = id_of_index conjecture_i  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
60  | 
val prem_ids = map id_of_index (prems_i upto prems_i + num_prems - 1)  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
61  | 
val fact_ids' =  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
62  | 
map_filter (fn j =>  | 
| 75367 | 63  | 
let val ((i, _), _) = nth assms j in  | 
| 
75299
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
64  | 
try (apsnd (nth xfacts)) (id_of_index j, i - facts_i)  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
65  | 
end) used_assm_js  | 
| 75367 | 66  | 
val helper_ids' =  | 
67  | 
map_filter (fn ((i, _), thm) => if i >= helpers_i then SOME (i, thm) else NONE) used_assms  | 
|
| 
75299
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
69  | 
val fact_helper_ts =  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
70  | 
map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, Thm.prop_of th)) helper_ids' @  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
71  | 
map (fn (_, ((s, _), th)) => (s, Thm.prop_of th)) fact_ids'  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
72  | 
val fact_helper_ids' =  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
73  | 
map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids' @ map (apsnd (fst o fst)) fact_ids'  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
74  | 
in  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
75  | 
    {outcome = NONE, fact_ids = SOME fact_ids',
 | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
76  | 
atp_proof = fn () => atp_proof_of_veriT_proof ctxt ll_defs rewrite_rules prems concl  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
77  | 
fact_helper_ts prem_ids conjecture_id fact_helper_ids' steps}  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
78  | 
end  | 
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
79  | 
|
| 
 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents:  
diff
changeset
 | 
80  | 
end;  |