author | blanchet |
Thu, 02 Oct 2014 17:39:38 +0200 | |
changeset 58515 | 6cf55e935a9d |
parent 58492 | e3ee0cf5cf93 |
child 72458 | b44e894796d5 |
permissions | -rw-r--r-- |
58061 | 1 |
(* Title: HOL/Tools/SMT/verit_isar.ML |
57704 | 2 |
Author: Mathias Fleury, TU Muenchen |
3 |
Author: Jasmin Blanchette, TU Muenchen |
|
4 |
||
5 |
VeriT proofs as generic ATP proofs for Isar proof reconstruction. |
|
6 |
*) |
|
7 |
||
8 |
signature VERIT_ISAR = |
|
9 |
sig |
|
10 |
type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step |
|
11 |
val atp_proof_of_veriT_proof: Proof.context -> term list -> thm list -> term list -> term -> |
|
12 |
(string * term) list -> int list -> int -> (int * string) list -> VeriT_Proof.veriT_step list -> |
|
13 |
(term, string) ATP_Proof.atp_step list |
|
14 |
end; |
|
15 |
||
16 |
structure VeriT_Isar: VERIT_ISAR = |
|
17 |
struct |
|
18 |
||
19 |
open ATP_Util |
|
20 |
open ATP_Problem |
|
21 |
open ATP_Proof |
|
22 |
open ATP_Proof_Reconstruct |
|
58482
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58064
diff
changeset
|
23 |
open SMTLIB_Interface |
58061 | 24 |
open SMTLIB_Isar |
57704 | 25 |
open VeriT_Proof |
26 |
||
27 |
fun atp_proof_of_veriT_proof ctxt ll_defs rewrite_rules hyp_ts concl_t fact_helper_ts prem_ids |
|
58492
e3ee0cf5cf93
repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents:
58489
diff
changeset
|
28 |
conjecture_id fact_helper_ids = |
57704 | 29 |
let |
30 |
fun steps_of (VeriT_Proof.VeriT_Step {id, rule, prems, concl, ...}) = |
|
31 |
let |
|
58064
e9ab6f4c650b
keep skolems as is -- otherwise this breaks proof steps that refer to variables that were 'obtain'ed in the outer Isar proof
blanchet
parents:
58061
diff
changeset
|
32 |
val concl' = postprocess_step_conclusion ctxt rewrite_rules ll_defs concl |
e9ab6f4c650b
keep skolems as is -- otherwise this breaks proof steps that refer to variables that were 'obtain'ed in the outer Isar proof
blanchet
parents:
58061
diff
changeset
|
33 |
fun standard_step role = ((id, []), role, concl', rule, map (rpair []) prems) |
57704 | 34 |
in |
57705 | 35 |
if rule = veriT_input_rule then |
58482
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58064
diff
changeset
|
36 |
let |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58064
diff
changeset
|
37 |
val id_num = the (Int.fromString (unprefix assert_prefix id)) |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58064
diff
changeset
|
38 |
val ss = the_list (AList.lookup (op =) fact_helper_ids id_num) |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58064
diff
changeset
|
39 |
in |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58064
diff
changeset
|
40 |
(case distinguish_conjecture_and_hypothesis ss id_num conjecture_id prem_ids |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58064
diff
changeset
|
41 |
fact_helper_ts hyp_ts concl_t of |
57730
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
42 |
NONE => [] |
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
43 |
| SOME (role0, concl00) => |
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
44 |
let |
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
45 |
val name0 = (id ^ "a", ss) |
58064
e9ab6f4c650b
keep skolems as is -- otherwise this breaks proof steps that refer to variables that were 'obtain'ed in the outer Isar proof
blanchet
parents:
58061
diff
changeset
|
46 |
val concl0 = unskolemize_names ctxt concl00 |
57730
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
47 |
in |
57745 | 48 |
[(name0, role0, concl0, rule, []), |
57730
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
49 |
((id, []), Plain, concl', veriT_rewrite_rule, |
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
50 |
name0 :: normalizing_prems ctxt concl0)] |
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
51 |
end) |
d39815cdb7ba
remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents:
57714
diff
changeset
|
52 |
end |
57704 | 53 |
else |
58515
6cf55e935a9d
avoid duplicate 'obtain' in veriT Isar proofs, by removing dubious condition
blanchet
parents:
58492
diff
changeset
|
54 |
[standard_step (if null prems then Lemma else Plain)] |
57704 | 55 |
end |
56 |
in |
|
58492
e3ee0cf5cf93
repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents:
58489
diff
changeset
|
57 |
maps steps_of |
57704 | 58 |
end |
59 |
||
60 |
end; |