| author | blanchet | 
| Mon, 02 Jul 2018 10:02:44 +0200 | |
| changeset 68563 | 05fb05f94686 | 
| parent 67022 | 49309fe530fd | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 57706 | 1  | 
(* Title: HOL/Tools/ATP/atp_satallax.ML  | 
2  | 
Author: Mathias Fleury, ENS Rennes  | 
|
3  | 
Author: Jasmin Blanchette, TU Muenchen  | 
|
4  | 
||
5  | 
Satallax proof parser and transformation for Sledgehammer.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature ATP_SATALLAX =  | 
|
9  | 
sig  | 
|
10  | 
val atp_proof_of_tstplike_proof : string -> string ATP_Proof.atp_problem -> string ->  | 
|
11  | 
string ATP_Proof.atp_proof  | 
|
12  | 
end;  | 
|
13  | 
||
14  | 
structure ATP_Satallax : ATP_SATALLAX =  | 
|
15  | 
struct  | 
|
16  | 
||
17  | 
open ATP_Proof  | 
|
18  | 
open ATP_Util  | 
|
19  | 
open ATP_Problem  | 
|
20  | 
||
21  | 
(*Undocumented format:  | 
|
22  | 
thf (number, plain | Axiom | ..., inference(rule, [status(thm), assumptions ([hypotheses list]),  | 
|
23  | 
detailed_rule(discharge,used_hypothese_list) *], used_hypotheses_list, premises))  | 
|
24  | 
(seen by tab_mat)  | 
|
25  | 
||
26  | 
Also seen -- but we can ignore it:  | 
|
27  | 
"tab_inh (a) __11." meaning that the type a is inhabited usueful of variable eigen__11,  | 
|
28  | 
*)  | 
|
29  | 
fun parse_satallax_tstp_information x =  | 
|
30  | 
((Symbol.scan_ascii_id || ($$ "$" |-- Symbol.scan_ascii_id))  | 
|
| 58325 | 31  | 
  -- Scan.option ( $$ "("
 | 
32  | 
|-- (Scan.option (Symbol.scan_ascii_id --| $$ ",")  | 
|
33  | 
-- ((Scan.option (($$ "[" |-- (Scan.option ((scan_general_id  | 
|
| 57716 | 34  | 
-- Scan.repeat ($$ "," |-- scan_general_id)) >> op ::) --| $$ "]"))  | 
| 58325 | 35  | 
|| (scan_general_id) >> (fn x => SOME [x]))  | 
36  | 
>> (fn (SOME x) => x | NONE => NONE)) --| $$ ")"))  | 
|
37  | 
|| (skip_term >> K (NONE, NONE)))) x  | 
|
| 57706 | 38  | 
|
39  | 
fun parse_prem x =  | 
|
40  | 
((Symbol.scan_ascii_id || scan_general_id) --| Scan.option ($$ ":" -- skip_term)) x  | 
|
41  | 
||
42  | 
fun parse_prems x =  | 
|
43  | 
(($$ "[" |-- parse_prem -- Scan.repeat ($$ "," |-- parse_prem) --| $$ "]")  | 
|
| 57716 | 44  | 
>> op ::) x  | 
| 57706 | 45  | 
|
46  | 
fun parse_tstp_satallax_extra_arguments x =  | 
|
47  | 
  ($$ "," |-- scan_general_id -- (($$ "(" |-- Symbol.scan_ascii_id --| $$ "," )
 | 
|
48  | 
-- ($$ "[" |-- Scan.option ((parse_satallax_tstp_information  | 
|
| 57716 | 49  | 
-- Scan.repeat ($$ "," |-- parse_satallax_tstp_information)) >> op ::)  | 
| 57706 | 50  | 
--| $$ "]") --  | 
51  | 
(Scan.optional ($$ "," |-- parse_prems) [] -- Scan.optional ($$ "," |-- parse_prems) []  | 
|
52  | 
>> (fn (x, []) => x | (_, x) => x))  | 
|
53  | 
--| $$ ")")) x  | 
|
54  | 
||
55  | 
val dummy_satallax_step = ((("~1", "tab_inh"), AAtom (ATerm(("False", []), []))), NONE)
 | 
|
56  | 
fun parse_tstp_thf0_satallax_line x =  | 
|
57  | 
(((Scan.this_string tptp_thf  | 
|
58  | 
  -- $$ "(") |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ ","
 | 
|
| 
67022
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
63692 
diff
changeset
 | 
59  | 
-- parse_thf_formula -- Scan.option parse_tstp_satallax_extra_arguments --| $$ ")" --| $$ ".")  | 
| 57706 | 60  | 
|| (Scan.this_string "tab_inh" |-- skip_term --| $$ ".")  | 
61  | 
>> K dummy_satallax_step) x  | 
|
62  | 
||
63  | 
datatype satallax_step = Satallax_Step of {
 | 
|
64  | 
id: string,  | 
|
65  | 
role: string,  | 
|
66  | 
theorem: (string, string, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, string)  | 
|
67  | 
ATP_Problem.atp_formula,  | 
|
68  | 
assumptions: string list,  | 
|
69  | 
rule: string,  | 
|
70  | 
used_assumptions: string list,  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
71  | 
detailed_eigen: string,  | 
| 57706 | 72  | 
generated_goal_assumptions: (string * string list) list}  | 
73  | 
||
74  | 
fun mk_satallax_step id role theorem assumptions rule used_assumptions  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
75  | 
generated_goal_assumptions detailed_eigen =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
76  | 
  Satallax_Step {id = id, role = role, theorem = theorem, assumptions = assumptions, rule = rule,
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
77  | 
used_assumptions = used_assumptions, generated_goal_assumptions = generated_goal_assumptions,  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
78  | 
detailed_eigen = detailed_eigen}  | 
| 57706 | 79  | 
|
80  | 
fun get_assumptions (("assumptions", SOME (_ , assumptions)) :: _) =
 | 
|
81  | 
the_default [] assumptions  | 
|
82  | 
| get_assumptions (_ :: l) = get_assumptions l  | 
|
83  | 
| get_assumptions [] = []  | 
|
84  | 
||
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
85  | 
fun get_detailled_eigen ((_, SOME (SOME "eigenvar" , var)) :: _) =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
86  | 
hd (the_default [""] var)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
87  | 
| get_detailled_eigen (_ :: l) = get_detailled_eigen l  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
88  | 
| get_detailled_eigen [] = ""  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
89  | 
|
| 57706 | 90  | 
fun seperate_dependices dependencies =  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
91  | 
let  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
92  | 
fun sep_dep [] used_assumptions new_goals generated_assumptions _ =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
93  | 
(used_assumptions, new_goals, generated_assumptions)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
94  | 
| sep_dep (x :: l) used_assumptions new_goals generated_assumptions state =  | 
| 58325 | 95  | 
if hd (raw_explode x) = "h" orelse Int.fromString x = NONE then  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
96  | 
if state = 0 then  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
97  | 
sep_dep l (x :: used_assumptions) new_goals generated_assumptions state  | 
| 57706 | 98  | 
else  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
99  | 
sep_dep l used_assumptions new_goals (x :: generated_assumptions) 3  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
100  | 
else  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
101  | 
if state = 1 orelse state = 0 then  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
102  | 
sep_dep l used_assumptions (x :: new_goals) generated_assumptions 1  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
103  | 
else  | 
| 58043 | 104  | 
            raise Fail ("incorrect Satallax proof: " ^ @{make_string} l)
 | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
105  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
106  | 
sep_dep dependencies [] [] [] 0  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
107  | 
end  | 
| 57706 | 108  | 
|
109  | 
fun create_grouped_goal_assumption rule new_goals generated_assumptions =  | 
|
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
110  | 
let  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
111  | 
val number_of_new_goals = length new_goals  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
112  | 
val number_of_new_assms = length generated_assumptions  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
113  | 
in  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
114  | 
if number_of_new_goals = number_of_new_assms then  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
115  | 
new_goals ~~ (map single generated_assumptions)  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
116  | 
else if 2 * number_of_new_goals = number_of_new_assms then  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
117  | 
let  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
118  | 
fun group_by_pair (new_goal :: new_goals) (assumpt1 :: assumpt2 :: generated_assumptions) =  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
119  | 
(new_goal, [assumpt1, assumpt2]) :: group_by_pair new_goals generated_assumptions  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
120  | 
| group_by_pair [] [] = []  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
121  | 
in  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
122  | 
group_by_pair new_goals generated_assumptions  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
123  | 
end  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
124  | 
else  | 
| 63692 | 125  | 
      raise Fail ("the rule " ^ rule ^" is not supported in the reconstruction")
 | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
126  | 
end  | 
| 57706 | 127  | 
fun to_satallax_step (((id, role), formula), (SOME (_,((rule, l), used_rules)))) =  | 
128  | 
let  | 
|
129  | 
val (used_assumptions, new_goals, generated_assumptions) = seperate_dependices used_rules  | 
|
130  | 
in  | 
|
131  | 
mk_satallax_step id role formula (get_assumptions (the_default [] l)) rule used_assumptions  | 
|
132  | 
(create_grouped_goal_assumption rule new_goals generated_assumptions)  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
133  | 
(get_detailled_eigen (the_default [] l))  | 
| 57706 | 134  | 
end  | 
135  | 
| to_satallax_step (((id, role), formula), NONE) =  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
136  | 
mk_satallax_step id role formula [] "assumption" [] [] ""  | 
| 57706 | 137  | 
|
138  | 
fun is_assumption (Satallax_Step {role, ...}) = role = "assumption" orelse role = "axiom" orelse
 | 
|
139  | 
role = "negated_conjecture" orelse role = "conjecture"  | 
|
140  | 
||
141  | 
fun seperate_assumptions_and_steps l =  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
142  | 
let  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
143  | 
fun seperate_assumption [] l l' = (l, l')  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
144  | 
| seperate_assumption (step :: steps) l l' =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
145  | 
if is_assumption step then  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
146  | 
seperate_assumption steps (step :: l) l'  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
147  | 
else  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
148  | 
seperate_assumption steps l (step :: l')  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
149  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
150  | 
seperate_assumption l [] []  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
151  | 
end  | 
| 57706 | 152  | 
|
153  | 
datatype satallax_proof_graph =  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
154  | 
  Linear_Part of {node: satallax_step, succs: satallax_proof_graph list} |
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
155  | 
  Tree_Part of {node: satallax_step, deps: satallax_proof_graph list}
 | 
| 57706 | 156  | 
|
157  | 
fun find_proof_step ((x as Satallax_Step {id, ...}) :: l) h =
 | 
|
158  | 
if h = id then x else find_proof_step l h  | 
|
| 58043 | 159  | 
  | find_proof_step [] h = raise Fail ("not_found: " ^ @{make_string} h ^ " (probably a parsing \
 | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
160  | 
\error)")  | 
| 57706 | 161  | 
|
162  | 
fun remove_not_not (x as ATerm ((op1, _), [ATerm ((op2, _), [th])])) =  | 
|
163  | 
if op1 = op2 andalso op1 = tptp_not then th else x  | 
|
164  | 
| remove_not_not th = th  | 
|
165  | 
||
166  | 
fun tptp_term_equal (ATerm((op1, _), l1), ATerm((op2, _), l2)) = op1 = op2 andalso  | 
|
167  | 
fold2 (fn t1 => fn t2 => fn c => c andalso t1 = t2) l1 l2 true  | 
|
168  | 
| tptp_term_equal (_, _) = false  | 
|
169  | 
||
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
170  | 
val dummy_true_aterm = ATerm (("$true", [dummy_atype]), [])
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
171  | 
|
| 57706 | 172  | 
fun find_assumptions_to_inline ths (assm :: assms) to_inline no_inline =  | 
173  | 
(case List.find (curry (op =) assm o fst) no_inline of  | 
|
174  | 
SOME _ => find_assumptions_to_inline ths assms to_inline no_inline  | 
|
175  | 
| NONE =>  | 
|
176  | 
(case List.find (curry (op =) assm o fst) to_inline of  | 
|
177  | 
NONE => find_assumptions_to_inline ths assms to_inline no_inline  | 
|
178  | 
| SOME (_, th) =>  | 
|
179  | 
let  | 
|
| 58325 | 180  | 
val simplified_ths_with_inlined_asms =  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
181  | 
(case List.partition (curry tptp_term_equal th o remove_not_not) ths of  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
182  | 
([], ths) => ATerm ((tptp_not, [dummy_atype]), [th]) :: ths  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
183  | 
| (_, _) => [])  | 
| 57706 | 184  | 
in  | 
| 58325 | 185  | 
find_assumptions_to_inline simplified_ths_with_inlined_asms assms to_inline no_inline  | 
| 57706 | 186  | 
end))  | 
187  | 
| find_assumptions_to_inline ths [] _ _ = ths  | 
|
188  | 
||
189  | 
fun inline_if_needed_and_simplify th assms to_inline no_inline =  | 
|
| 58325 | 190  | 
(case find_assumptions_to_inline [th] assms to_inline no_inline of  | 
191  | 
[] => dummy_true_aterm  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
192  | 
| l => foldl1 (fn (a, b) =>  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
193  | 
(case b of  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
194  | 
      ATerm (("$false", _), _) => a
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
195  | 
    | ATerm (("~", _ ), [ATerm (("$true", _), _)]) => a
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
196  | 
| _ => ATerm ((tptp_or, [dummy_atype]), [a, b]))) l)  | 
| 57706 | 197  | 
|
198  | 
fun get_conclusion (Satallax_Step {theorem = AAtom theorem, ...}) = theorem
 | 
|
199  | 
||
| 58325 | 200  | 
fun add_assumptions new_used_assumptions (Satallax_Step {id, role, theorem, assumptions,
 | 
201  | 
rule, generated_goal_assumptions, used_assumptions, detailed_eigen}) =  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
202  | 
mk_satallax_step id role theorem assumptions rule (new_used_assumptions @ used_assumptions)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
203  | 
generated_goal_assumptions detailed_eigen  | 
| 57706 | 204  | 
|
205  | 
fun set_rule new_rule (Satallax_Step {id, role, theorem, assumptions,
 | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
206  | 
generated_goal_assumptions, used_assumptions, detailed_eigen, ...}) =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
207  | 
mk_satallax_step id role theorem assumptions new_rule used_assumptions  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
208  | 
generated_goal_assumptions detailed_eigen  | 
| 57706 | 209  | 
|
| 58325 | 210  | 
fun add_detailled_eigen eigen (Satallax_Step {id, role, theorem, assumptions,
 | 
211  | 
rule, generated_goal_assumptions, used_assumptions, detailed_eigen}) =  | 
|
212  | 
mk_satallax_step id role theorem assumptions rule used_assumptions  | 
|
213  | 
generated_goal_assumptions (if detailed_eigen <> "" then detailed_eigen else eigen)  | 
|
214  | 
||
| 57706 | 215  | 
fun transform_inline_assumption hypotheses_step proof_sketch =  | 
216  | 
let  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
217  | 
    fun inline_in_step (Linear_Part {node as Satallax_Step {generated_goal_assumptions,
 | 
| 58325 | 218  | 
used_assumptions, rule, detailed_eigen, ...}, succs}) =  | 
| 57706 | 219  | 
if generated_goal_assumptions = [] then  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
220  | 
          Linear_Part {node = node, succs = []}
 | 
| 57706 | 221  | 
else  | 
222  | 
let  | 
|
223  | 
(*one singel goal is created, two hypothesis can be created, for the "and" rule:  | 
|
224  | 
(A /\ B) create two hypotheses A, and B.*)  | 
|
225  | 
fun set_hypotheses_as_goal [hypothesis] succs =  | 
|
| 58325 | 226  | 
                Linear_Part {node = add_detailled_eigen detailed_eigen
 | 
227  | 
(set_rule rule (add_assumptions used_assumptions  | 
|
228  | 
(find_proof_step hypotheses_step hypothesis))),  | 
|
| 57706 | 229  | 
succs = map inline_in_step succs}  | 
230  | 
| set_hypotheses_as_goal (hypothesis :: new_goal_hypotheses) succs =  | 
|
| 58325 | 231  | 
                Linear_Part {node = set_rule rule (add_assumptions used_assumptions
 | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
232  | 
(find_proof_step hypotheses_step hypothesis)),  | 
| 57706 | 233  | 
succs = [set_hypotheses_as_goal new_goal_hypotheses succs]}  | 
234  | 
in  | 
|
235  | 
set_hypotheses_as_goal (snd (hd generated_goal_assumptions)) succs  | 
|
236  | 
end  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
237  | 
      | inline_in_step (Tree_Part {node = node, deps}) =
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
238  | 
        Tree_Part {node = node, deps = map inline_in_step deps}
 | 
| 57706 | 239  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
240  | 
    fun inline_contradictory_assumptions (Linear_Part {node as Satallax_Step{id, theorem, ...},
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
241  | 
succs}) (to_inline, no_inline) =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
242  | 
let  | 
| 58325 | 243  | 
val (succs, inliner) = fold_map inline_contradictory_assumptions succs  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
244  | 
(to_inline, (id, theorem) :: no_inline)  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
245  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
246  | 
        (Linear_Part {node = node, succs = succs}, inliner)
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
247  | 
end  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
248  | 
    | inline_contradictory_assumptions (Tree_Part {node = Satallax_Step {id, role,
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
249  | 
theorem = AAtom theorem, assumptions, rule, generated_goal_assumptions,  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
250  | 
used_assumptions, detailed_eigen}, deps}) (to_inline, no_inline) =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
251  | 
let  | 
| 58325 | 252  | 
val (dep', (to_inline', no_inline')) = fold_map inline_contradictory_assumptions deps  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
253  | 
(to_inline, no_inline)  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
254  | 
val to_inline'' =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
255  | 
map (fn s => (s, get_conclusion (find_proof_step hypotheses_step s)))  | 
| 58325 | 256  | 
(filter (fn s => (nth_string s 0 = "h") andalso List.find (curry (op =) s o fst)  | 
257  | 
no_inline' = NONE) (used_assumptions @ (map snd generated_goal_assumptions |> flat)))  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
258  | 
@ to_inline'  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
259  | 
        val node' = Satallax_Step {id = id, role = role, theorem =
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
260  | 
AAtom (inline_if_needed_and_simplify theorem assumptions to_inline'' no_inline'),  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
261  | 
assumptions = assumptions, rule = rule,  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
262  | 
generated_goal_assumptions = generated_goal_assumptions, detailed_eigen = detailed_eigen,  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
263  | 
used_assumptions =  | 
| 58325 | 264  | 
filter (fn s => List.find (curry (op =) s o fst) to_inline'' = NONE)  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
265  | 
used_assumptions}  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
266  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
267  | 
        (Tree_Part {node = node', deps = dep'}, (to_inline'', no_inline'))
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
268  | 
end  | 
| 57706 | 269  | 
in  | 
270  | 
fst (inline_contradictory_assumptions (inline_in_step proof_sketch) ([], []))  | 
|
271  | 
end  | 
|
272  | 
||
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
273  | 
fun correct_dependencies (Linear_Part {node, succs}) =
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
274  | 
    Linear_Part {node = node, succs = map correct_dependencies succs}
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
275  | 
  | correct_dependencies (Tree_Part {node, deps}) =
 | 
| 57706 | 276  | 
let  | 
277  | 
val new_used_assumptions =  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
278  | 
        map (fn Linear_Part {node = (Satallax_Step{id, ...}), ...} => id
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
279  | 
              | Tree_Part {node = (Satallax_Step{id, ...}), ...} => id) deps
 | 
| 57706 | 280  | 
in  | 
| 58325 | 281  | 
      Tree_Part {node = add_assumptions new_used_assumptions node,
 | 
| 57706 | 282  | 
deps = map correct_dependencies deps}  | 
283  | 
end  | 
|
284  | 
||
285  | 
fun create_satallax_proof_graph (hypotheses_step, proof_sketch) =  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
286  | 
let  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
287  | 
    fun create_step (step as Satallax_Step {generated_goal_assumptions, ...}) =
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
288  | 
      Linear_Part {node = step,
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
289  | 
succs = map (create_step o (find_proof_step (hypotheses_step @ proof_sketch)))  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
290  | 
(map fst generated_goal_assumptions)}  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
291  | 
    fun reverted_discharged_steps is_branched (Linear_Part {node as
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
292  | 
          Satallax_Step {generated_goal_assumptions, ...}, succs}) =
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
293  | 
if is_branched orelse length generated_goal_assumptions > 1 then  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
294  | 
          Tree_Part {node = node, deps = map (reverted_discharged_steps true) succs}
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
295  | 
else  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
296  | 
          Linear_Part {node = node, succs = map (reverted_discharged_steps is_branched) succs}
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
297  | 
val proof_with_correct_sense =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
298  | 
correct_dependencies (reverted_discharged_steps false  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
299  | 
(create_step (find_proof_step proof_sketch "0" )))  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
300  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
301  | 
transform_inline_assumption hypotheses_step proof_with_correct_sense  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
302  | 
end  | 
| 57706 | 303  | 
|
| 58325 | 304  | 
val satallax_known_rules = ["eq_ind", "eq_trans2", "eq_trans3", "eq_neg_neg_id", "eq_true",  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
305  | 
"eq_and_nor", "eq_or_nand", "eq_or_imp", "eq_and_imp", "eq_imp_or", "eq_iff", "eq_sym_eq",  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
306  | 
"eq_forall_nexists", "eq_exists_nforall", "eq_leib1", "eq_leib2", "eq_leib3", "eq_leib4",  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
307  | 
"eq_eta", "SinhE", "eq_forall_Seps", "eq_SPi_Seps", "eq_exists_Seps"]  | 
| 58325 | 308  | 
val is_special_satallax_rule = member (op =) satallax_known_rules  | 
| 57706 | 309  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
310  | 
fun terms_to_upper_case var (AAbs (((var', ty), b), ts)) =  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
311  | 
let  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
312  | 
val bdy = terms_to_upper_case var b  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
313  | 
val ts' = map (terms_to_upper_case var) ts  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
314  | 
in  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
315  | 
AAbs (((((var = var' ? String.implode o (map Char.toUpper) o String.explode) var'), ty),  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
316  | 
bdy), ts')  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
317  | 
end  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
318  | 
| terms_to_upper_case var (ATerm ((var', ty), ts)) =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
319  | 
ATerm ((((var = var' ? String.implode o (map Char.toUpper) o String.explode) var'),  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
320  | 
ty), map (terms_to_upper_case var) ts)  | 
| 57706 | 321  | 
|
| 58325 | 322  | 
fun add_quantifier_in_tree_part var_rule_to_add assumption_to_add  | 
323  | 
      (Linear_Part {node as Satallax_Step {detailed_eigen, rule, ...} , succs}) =
 | 
|
324  | 
    Linear_Part {node = node, succs = map (add_quantifier_in_tree_part
 | 
|
325  | 
((detailed_eigen <> "" ? cons (detailed_eigen, rule)) var_rule_to_add) assumption_to_add)  | 
|
326  | 
succs}  | 
|
327  | 
| add_quantifier_in_tree_part var_rule_to_add assumption_to_add  | 
|
328  | 
      (Tree_Part {node = Satallax_Step {rule, detailed_eigen, id, role, theorem = AAtom th,
 | 
|
329  | 
assumptions, used_assumptions, generated_goal_assumptions}, deps}) =  | 
|
| 57706 | 330  | 
let  | 
| 58325 | 331  | 
val theorem' = fold (fn v => fn body => terms_to_upper_case (fst v) body) var_rule_to_add th  | 
332  | 
fun add_quantified_var (var, rule) = fn body =>  | 
|
333  | 
let  | 
|
334  | 
val quant = if rule = "tab_ex" then tptp_ho_exists else tptp_ho_forall  | 
|
335  | 
val upperVar = (String.implode o (map Char.toUpper) o String.explode) var  | 
|
336  | 
val quant_bdy = if rule = "tab_ex"  | 
|
337  | 
then ATerm ((quant, []), [AAbs (((upperVar, dummy_atype), body), []) ]) else body  | 
|
338  | 
in  | 
|
339  | 
quant_bdy  | 
|
340  | 
end  | 
|
341  | 
val theorem'' = fold add_quantified_var var_rule_to_add theorem'  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
342  | 
val node' = mk_satallax_step id role (AAtom theorem'') assumptions rule  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
343  | 
(used_assumptions @ (filter (curry (op <=) (the (Int.fromString id)) o size)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
344  | 
assumption_to_add)) generated_goal_assumptions detailed_eigen  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
345  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
346  | 
if detailed_eigen <> "" then  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
347  | 
        Tree_Part {node = node',
 | 
| 58325 | 348  | 
deps = map (add_quantifier_in_tree_part ((detailed_eigen, rule) :: var_rule_to_add)  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
349  | 
(used_assumptions @ assumption_to_add)) deps}  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
350  | 
else  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
351  | 
        Tree_Part {node = node',
 | 
| 58325 | 352  | 
deps = map (add_quantifier_in_tree_part var_rule_to_add assumption_to_add) deps}  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
353  | 
end  | 
| 57706 | 354  | 
|
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
355  | 
fun transform_to_standard_atp_step already_transformed proof =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
356  | 
let  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
357  | 
fun create_fact_step id =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
358  | 
((id, [id]), Axiom, AAtom (ATerm((id, []), [])), "", [])  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
359  | 
    fun transform_one_step already_transformed (Satallax_Step {id, theorem, used_assumptions, role,
 | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
360  | 
rule, ...}) =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
361  | 
let  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
362  | 
val role' = role_of_tptp_string role  | 
| 58412 | 363  | 
val new_transformed = filter  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
364  | 
(fn s => size s >=4 andalso not (is_special_satallax_rule s) andalso not  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
365  | 
(member (op =) already_transformed s)) used_assumptions  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
366  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
367  | 
(map create_fact_step new_transformed  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
368  | 
@ [((id, []), if role' = Unknown then Plain else role', theorem, rule,  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
369  | 
map (fn s => (s, [])) (filter_out is_special_satallax_rule used_assumptions))],  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
370  | 
new_transformed @ already_transformed)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
371  | 
end  | 
| 58325 | 372  | 
    fun transform_steps (Linear_Part {node, succs}) already_transformed =
 | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
373  | 
transform_one_step already_transformed node  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
374  | 
||> (fold_map transform_steps succs)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
375  | 
||> apfst flat  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
376  | 
|> (fn (a, (b, transformed)) => (a @ b, transformed))  | 
| 58325 | 377  | 
      | transform_steps (Tree_Part {node, deps}) already_transformed =
 | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
378  | 
fold_map transform_steps deps already_transformed  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
379  | 
|>> flat  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
380  | 
||> (fn transformed => transform_one_step transformed node)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
381  | 
|> (fn (a, (b, transformed)) => (a @ b, transformed))  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
382  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
383  | 
fst (transform_steps proof already_transformed)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
384  | 
end  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
385  | 
|
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
386  | 
fun remove_unused_dependency steps =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
387  | 
let  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
388  | 
fun find_all_ids [] = []  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
389  | 
| find_all_ids (((id, _), _, _, _, _) :: steps) = id :: find_all_ids steps  | 
| 
57710
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
390  | 
fun keep_only_used used_ids steps =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
391  | 
let  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
392  | 
fun remove_unused (((id, ids), role, theorem, rule, deps) :: steps) =  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
393  | 
(((id, ids), role, theorem, rule, filter (member (op =) used_ids o fst) deps) :: steps)  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
394  | 
| remove_unused [] = []  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
395  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
396  | 
remove_unused steps  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
397  | 
end  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
398  | 
in  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
399  | 
keep_only_used (find_all_ids steps) steps  | 
| 
 
323a57d7455c
imported patch satallax_skolemization_in_tree_part
 
fleury 
parents: 
57707 
diff
changeset
 | 
400  | 
end  | 
| 57706 | 401  | 
|
402  | 
fun parse_proof local_name problem =  | 
|
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
403  | 
strip_spaces_except_between_idents  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
404  | 
#> raw_explode  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
405  | 
#>  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
406  | 
(if local_name <> satallaxN then  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
407  | 
(Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ())  | 
| 61476 | 408  | 
(Scan.finite Symbol.stopper (Scan.repeats1 (parse_line local_name problem))))  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
409  | 
#> fst)  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
410  | 
else  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
411  | 
(Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ())  | 
| 
67022
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
63692 
diff
changeset
 | 
412  | 
(Scan.finite Symbol.stopper (Scan.repeat1 parse_tstp_thf0_satallax_line)))  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
413  | 
#> fst  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
414  | 
#> rev  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
415  | 
#> map to_satallax_step  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
416  | 
#> seperate_assumptions_and_steps  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
417  | 
#> create_satallax_proof_graph  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
418  | 
#> add_quantifier_in_tree_part [] []  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
419  | 
#> transform_to_standard_atp_step []  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57710 
diff
changeset
 | 
420  | 
#> remove_unused_dependency))  | 
| 57706 | 421  | 
|
422  | 
fun atp_proof_of_tstplike_proof _ _ "" = []  | 
|
423  | 
| atp_proof_of_tstplike_proof local_prover problem tstp =  | 
|
424  | 
(case core_of_agsyhol_proof tstp of  | 
|
425  | 
SOME facts => facts |> map (core_inference agsyhol_core_rule)  | 
|
426  | 
| NONE =>  | 
|
427  | 
tstp ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)  | 
|
428  | 
|> parse_proof local_prover problem  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58412 
diff
changeset
 | 
429  | 
|> local_prover = vampireN ? perhaps (try (sort (vampire_step_name_ord o apply2 #1)))  | 
| 57706 | 430  | 
|> local_prover = zipperpositionN ? rev)  | 
431  | 
||
432  | 
end;  |