src/HOL/Tools/SMT/verit_proof_parse.ML
author desharna
Fri, 11 Mar 2022 09:22:13 +0100
changeset 75274 e89709b80b6e
parent 72458 b44e894796d5
child 75365 83940294cc67
permissions -rw-r--r--
used more descriptive assert names in SMT-Lib output
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57716
diff changeset
     1
(*  Title:      HOL/Tools/SMT/verit_proof_parse.ML
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     2
    Author:     Mathias Fleury, TU Muenchen
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     4
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     5
VeriT proof parsing.
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     6
*)
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     7
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     8
signature VERIT_PROOF_PARSE =
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     9
sig
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    10
  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
58491
blanchet
parents: 58488
diff changeset
    11
  val parse_proof: SMT_Translate.replay_data ->
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    12
    ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57716
diff changeset
    13
    SMT_Solver.parsed_proof
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    14
end;
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    15
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    16
structure VeriT_Proof_Parse: VERIT_PROOF_PARSE =
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    17
struct
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    18
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    19
open ATP_Util
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    20
open ATP_Problem
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    21
open ATP_Proof
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    22
open ATP_Proof_Reconstruct
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    23
open VeriT_Isar
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69205
diff changeset
    24
open Verit_Proof
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    25
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69205
diff changeset
    26
fun add_used_asserts_in_step (Verit_Proof.VeriT_Step {prems, ...}) =
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 72458
diff changeset
    27
  union (op =) (map_filter (try (snd o SMTLIB_Interface.role_and_index_of_assert_name)) prems)
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    28
58491
blanchet
parents: 58488
diff changeset
    29
fun parse_proof
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57716
diff changeset
    30
    ({context = ctxt, typs, terms, ll_defs, rewrite_rules, assms} : SMT_Translate.replay_data)
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    31
    xfacts prems concl output =
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    32
  let
58492
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    33
    val num_ll_defs = length ll_defs
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    34
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    35
    val id_of_index = Integer.add num_ll_defs
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    36
    val index_of_id = Integer.add (~ num_ll_defs)
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    37
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    38
    fun step_of_assume j (_, th) =
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 72458
diff changeset
    39
      Verit_Proof.VeriT_Step
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 72458
diff changeset
    40
        {id = SMTLIB_Interface.assert_name_of_role_and_index SMT_Util.Axiom (id_of_index j),
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69205
diff changeset
    41
        rule = input_rule, prems = [], proof_ctxt = [], concl = Thm.prop_of th, fixes = []}
58492
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    42
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69205
diff changeset
    43
    val (actual_steps, _) = Verit_Proof.parse typs terms output ctxt
58492
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    44
    val used_assert_ids = fold add_used_asserts_in_step actual_steps []
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    45
    val used_assm_js =
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    46
      map_filter (fn id => let val i = index_of_id id in if i >= 0 then SOME i else NONE end)
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    47
        used_assert_ids
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    48
    val used_assms = map (nth assms) used_assm_js
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    49
    val assm_steps = map2 step_of_assume used_assm_js used_assms
58482
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    50
    val steps = assm_steps @ actual_steps
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    51
58492
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    52
    val conjecture_i = 0
58488
289d1c39968c correct indexing in the presence of lambda-lifting
blanchet
parents: 58482
diff changeset
    53
    val prems_i = conjecture_i + 1
58482
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    54
    val num_prems = length prems
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    55
    val facts_i = prems_i + num_prems
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    56
    val num_facts = length xfacts
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    57
    val helpers_i = facts_i + num_facts
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    58
58492
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    59
    val conjecture_id = id_of_index conjecture_i
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    60
    val prem_ids = map id_of_index (prems_i upto prems_i + num_prems - 1)
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    61
    val fact_ids' =
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    62
      map_filter (fn j =>
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    63
        let val (i, _) = nth assms j in
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    64
          try (apsnd (nth xfacts)) (id_of_index j, i - facts_i)
e3ee0cf5cf93 repaired index confusion -- in particular, carefully distinguish between 'assert indices' (monomorphised etc.) and 'assume indices'
blanchet
parents: 58491
diff changeset
    65
        end) used_assm_js
58482
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    66
    val helper_ids' = filter (fn (i, _) => i >= helpers_i) used_assms
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    67
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    68
    val fact_helper_ts =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58492
diff changeset
    69
      map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, Thm.prop_of th)) helper_ids' @
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58492
diff changeset
    70
      map (fn (_, ((s, _), th)) => (s, Thm.prop_of th)) fact_ids'
58482
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    71
    val fact_helper_ids' =
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    72
      map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids' @ map (apsnd (fst o fst)) fact_ids'
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    73
  in
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59582
diff changeset
    74
    {outcome = NONE, fact_ids = SOME fact_ids',
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    75
     atp_proof = fn () => atp_proof_of_veriT_proof ctxt ll_defs rewrite_rules prems concl
58482
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58061
diff changeset
    76
       fact_helper_ts prem_ids conjecture_id fact_helper_ids' steps}
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    77
  end
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    78
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    79
end;