src/HOL/Tools/SMT/smtlib_isar.ML
author haftmann
Fri, 27 Jun 2025 08:09:26 +0200
changeset 82775 61c39a9e5415
parent 82643 f1c14af17591
permissions -rw-r--r--
typo
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/smtlib_isar.ML
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
     3
    Author:     Mathias Fleury, ENS Rennes
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
General tools 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
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57768
diff changeset
     8
signature SMTLIB_ISAR =
57704
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
  val unlift_term: term list -> term -> term
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
    11
  val postprocess_step_conclusion: Proof.context -> thm list -> term list -> term -> term
57705
5da48dae7d03 Subproofs for the SMT solver veriT.
fleury
parents: 57704
diff changeset
    12
  val normalizing_prems : Proof.context -> term -> (string * string list) list
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    13
  val distinguish_conjecture_and_hypothesis : ''a list -> ''b -> ''b -> ''b list ->
58484
b4c0e2b00036 support hypotheses with schematics in Isar proofs
blanchet
parents: 58064
diff changeset
    14
    (''a * term) list -> term list -> term -> (ATP_Problem.atp_formula_role * term) option
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
    15
  val unskolemize_names: Proof.context -> term -> term
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    16
end;
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    17
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57768
diff changeset
    18
structure SMTLIB_Isar: SMTLIB_ISAR =
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    19
struct
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    20
57768
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57767
diff changeset
    21
open ATP_Util
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    22
open ATP_Problem
57768
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57767
diff changeset
    23
open ATP_Proof_Reconstruct
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    24
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    25
fun unlift_term ll_defs =
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    26
  let
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    27
    val lifted = map (ATP_Util.extract_lambda_def dest_Free o ATP_Util.hol_open_form I) ll_defs
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    28
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    29
    fun un_free (t as Free (s, _)) =
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    30
       (case AList.lookup (op =) lifted s of
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    31
         SOME t => un_term t
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    32
       | NONE => t)
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    33
     | un_free t = t
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    34
    and un_term t = map_aterms un_free t
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    35
  in un_term end
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    36
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
    37
(* Remove the "__" suffix for newly introduced variables (Skolems). It is not clear why "__" is
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
    38
   generated also for abstraction variables, but this is repaired here. *)
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
    39
fun unskolemize_names ctxt =
57708
4b52c1b319ce veriT changes for lifted terms, and ite_elim rules.
fleury
parents: 57705
diff changeset
    40
  Term.map_abs_vars (perhaps (try Name.dest_skolem))
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
    41
  #> Term.map_aterms (perhaps (try (fn Free (s, T) =>
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
    42
    Free (s |> not (Variable.is_fixed ctxt s) ? Name.dest_skolem, T))))
57708
4b52c1b319ce veriT changes for lifted terms, and ite_elim rules.
fleury
parents: 57705
diff changeset
    43
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
    44
fun postprocess_step_conclusion ctxt rewrite_rules ll_defs =
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
    45
  let val thy = Proof_Context.theory_of ctxt in
82643
f1c14af17591 prefer Simplifier over bootstrap-only Raw_Simplifier
haftmann
parents: 59970
diff changeset
    46
    Simplifier.rewrite_term thy rewrite_rules []
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 58484
diff changeset
    47
    #> Object_Logic.atomize_term ctxt
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
    48
    #> not (null ll_defs) ? unlift_term ll_defs
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
    49
    #> simplify_bool
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
    50
    #> unskolemize_names ctxt
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
    51
    #> HOLogic.mk_Trueprop
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
    52
  end
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    53
57705
5da48dae7d03 Subproofs for the SMT solver veriT.
fleury
parents: 57704
diff changeset
    54
fun normalizing_prems ctxt concl0 =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57768
diff changeset
    55
  SMT_Normalize.case_bool_entry :: SMT_Normalize.special_quant_table @
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57768
diff changeset
    56
  SMT_Normalize.abs_min_max_table
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    57
  |> map_filter (fn (c, th) =>
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    58
    if exists_Const (curry (op =) c o fst) concl0 then
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    59
      let val s = short_thm_name ctxt th in SOME (s, [s]) end
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    60
    else
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    61
      NONE)
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    62
57768
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57767
diff changeset
    63
fun distinguish_conjecture_and_hypothesis ss id conjecture_id prem_ids fact_helper_ts hyp_ts
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57767
diff changeset
    64
    concl_t =
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    65
  (case ss of
57730
d39815cdb7ba remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents: 57714
diff changeset
    66
    [s] => SOME (Axiom, the (AList.lookup (op =) fact_helper_ts s))
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    67
  | _ =>
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    68
    if id = conjecture_id then
57730
d39815cdb7ba remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents: 57714
diff changeset
    69
      SOME (Conjecture, concl_t)
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    70
    else
57730
d39815cdb7ba remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents: 57714
diff changeset
    71
     (case find_index (curry (op =) id) prem_ids of
d39815cdb7ba remove lambda-lifting related assumptions from generated Isar proofs
blanchet
parents: 57714
diff changeset
    72
       ~1 => NONE (* lambda-lifting definition *)
58484
b4c0e2b00036 support hypotheses with schematics in Isar proofs
blanchet
parents: 58064
diff changeset
    73
     | i => SOME (Hypothesis, close_form (nth hyp_ts i))))
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    74
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents:
diff changeset
    75
end;