src/HOL/Tools/SMT/verit_isar.ML
author blanchet
Thu, 02 Oct 2014 17:39:38 +0200
changeset 58515 6cf55e935a9d
parent 58492 e3ee0cf5cf93
child 72458 b44e894796d5
permissions -rw-r--r--
avoid duplicate 'obtain' in veriT Isar proofs, by removing dubious condition
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57768
diff changeset
     1
(*  Title:      HOL/Tools/SMT/verit_isar.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 proofs as generic ATP proofs for Isar proof reconstruction.
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_ISAR =
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
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    11
  val atp_proof_of_veriT_proof: Proof.context -> term list -> thm list -> term list -> term ->
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    12
    (string * term) list -> int list -> int -> (int * string) list -> VeriT_Proof.veriT_step list ->
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    13
    (term, string) ATP_Proof.atp_step list
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_Isar: VERIT_ISAR =
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
58482
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58064
diff changeset
    23
open SMTLIB_Interface
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57768
diff changeset
    24
open SMTLIB_Isar
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    25
open VeriT_Proof
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    26
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    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
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    29
  let
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    30
    fun steps_of (VeriT_Proof.VeriT_Step {id, rule, prems, concl, ...}) =
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    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
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    34
      in
57705
5da48dae7d03 Subproofs for the SMT solver veriT.
fleury
parents: 57704
diff changeset
    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
c65c116553b5 more rational unskolemizing of names
blanchet
parents: 57730
diff changeset
    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
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    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
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    55
      end
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    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
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    58
  end
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    59
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    60
end;