src/HOL/Tools/SMT/z3_replay_methods.ML
author wenzelm
Wed, 25 Jun 2025 16:35:25 +0200
changeset 82768 8f866fd6fae1
parent 74382 8d0294d877bd
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
     1
(*  Title:      HOL/Tools/SMT/z3_replay_methods.ML
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     4
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     5
Proof methods for replaying Z3 proofs.
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     6
*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     7
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
     8
signature Z3_REPLAY_METHODS =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     9
sig
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    10
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    11
  (*methods for Z3 proof rules*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    12
  type z3_method = Proof.context -> thm list -> term -> thm
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    13
  val true_axiom: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    14
  val mp: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
  val refl: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    16
  val symm: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
  val trans: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    18
  val cong: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
  val quant_intro: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    20
  val distrib: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    21
  val and_elim: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    22
  val not_or_elim: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    23
  val rewrite: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    24
  val rewrite_star: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
  val pull_quant: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    26
  val push_quant: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    27
  val elim_unused: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
  val dest_eq_res: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    29
  val quant_inst: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    30
  val lemma: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
  val unit_res: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    32
  val iff_true: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
  val iff_false: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    34
  val comm: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
  val def_axiom: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    36
  val apply_def: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    37
  val iff_oeq: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    38
  val nnf_pos: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    39
  val nnf_neg: z3_method
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
  val mp_oeq: z3_method
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    41
  val arith_th_lemma: z3_method
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    42
  val th_lemma: string -> z3_method
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    43
  val method_for: Z3_Proof.z3_rule -> z3_method
57229
blanchet
parents: 57220
diff changeset
    44
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    45
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    46
structure Z3_Replay_Methods: Z3_REPLAY_METHODS =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    47
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    48
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    49
type z3_method = Proof.context -> thm list -> term -> thm
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
(* utility functions *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    54
fun replay_error ctxt msg rule thms t =
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    55
  SMT_Replay_Methods.replay_error ctxt msg (Z3_Proof.string_of_rule rule) thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    56
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    57
fun replay_rule_error ctxt rule thms t =
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
    58
  SMT_Replay_Methods.replay_rule_error "Z3" ctxt (Z3_Proof.string_of_rule rule) thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    60
fun trace_goal ctxt rule thms t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    61
  SMT_Replay_Methods.trace_goal ctxt (Z3_Proof.string_of_rule rule) thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
    63
fun dest_prop (Const (\<^const_name>\<open>Trueprop\<close>, _) $ t) = t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    64
  | dest_prop t = t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
fun dest_thm thm = dest_prop (Thm.concl_of thm)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    67
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
    68
val try_provers = SMT_Replay_Methods.try_provers "Z3"
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
    69
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    70
fun prop_tac ctxt prems =
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61466
diff changeset
    71
  Method.insert_tac ctxt prems
56816
2f3756ccba41 use internal proof-producing SAT solver for more efficient SMT proof replay
boehmes
parents: 56090
diff changeset
    72
  THEN' SUBGOAL (fn (prop, i) =>
2f3756ccba41 use internal proof-producing SAT solver for more efficient SMT proof replay
boehmes
parents: 56090
diff changeset
    73
    if Term.size_of_term prop > 100 then SAT.satx_tac ctxt i
2f3756ccba41 use internal proof-producing SAT solver for more efficient SMT proof replay
boehmes
parents: 56090
diff changeset
    74
    else (Classical.fast_tac ctxt ORELSE' Clasimp.force_tac ctxt) i)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    75
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    76
fun quant_tac ctxt = Blast.blast_tac ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    77
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    78
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    79
fun apply_rule ctxt t =
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    80
  (case Z3_Replay_Rules.apply ctxt (SMT_Replay_Methods.certify_prop ctxt t) of
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    81
    SOME thm => thm
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    82
  | NONE => raise Fail "apply_rule")
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    83
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    84
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    85
(*theory-lemma*)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    86
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    87
fun arith_th_lemma_tac ctxt prems =
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    88
  Method.insert_tac ctxt prems
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    89
  THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms z3div_def z3mod_def})
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    90
  THEN' Arith_Data.arith_tac ctxt
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    91
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    92
fun arith_th_lemma ctxt thms t =
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    93
  SMT_Replay_Methods.prove_abstract ctxt thms t arith_th_lemma_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    94
    fold_map (SMT_Replay_Methods.abstract_arith ctxt o dest_thm) thms ##>>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    95
    SMT_Replay_Methods.abstract_arith ctxt (dest_prop t))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    96
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    97
fun th_lemma name ctxt thms =
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    98
  (case Symtab.lookup (SMT_Replay_Methods.get_th_lemma_method ctxt) name of
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
    99
    SOME method => method ctxt thms
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   100
  | NONE => replay_error ctxt "Bad theory" (Z3_Proof.Th_Lemma name) thms)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   101
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   102
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
(* truth axiom *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   104
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
fun true_axiom _ _ _ = @{thm TrueI}
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   108
(* modus ponens *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   109
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   110
fun mp _ [p, p_eq_q] _ = SMT_Replay_Methods.discharge 1 [p_eq_q, p] iffD1
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   111
  | mp ctxt thms t = replay_rule_error ctxt Z3_Proof.Modus_Ponens thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   112
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   113
val mp_oeq = mp
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   114
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   115
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   116
(* reflexivity *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   117
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   118
fun refl ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t @{thm refl}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   119
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   120
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   121
(* symmetry *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   122
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   123
fun symm _ [thm] _ = thm RS @{thm sym}
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   124
  | symm ctxt thms t = replay_rule_error ctxt Z3_Proof.Reflexivity thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   125
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   126
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   127
(* transitivity *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   128
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   129
fun trans _ [thm1, thm2] _ = thm1 RSN (1, thm2 RSN (2, @{thm trans}))
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   130
  | trans ctxt thms t = replay_rule_error ctxt Z3_Proof.Transitivity thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   131
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   132
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   133
(* congruence *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   134
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
   135
fun cong ctxt thms = try_provers ctxt
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   136
    (Z3_Proof.string_of_rule Z3_Proof.Monotonicity) [
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   137
  ("basic", SMT_Replay_Methods.cong_basic ctxt thms),
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   138
  ("full", SMT_Replay_Methods.cong_full ctxt thms)] thms
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   139
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   140
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
(* quantifier introduction *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   142
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   143
val quant_intro_rules = @{lemma
67091
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   144
  "(\<And>x. P x = Q x) ==> (\<forall>x. P x) = (\<forall>x. Q x)"
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   145
  "(\<And>x. P x = Q x) ==> (\<exists>x. P x) = (\<exists>x. Q x)"
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   146
  "(!!x. (\<not> P x) = Q x) \<Longrightarrow> (\<not>(\<exists>x. P x)) = (\<forall>x. Q x)"
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   147
  "(\<And>x. (\<not> P x) = Q x) ==> (\<not>(\<forall>x. P x)) = (\<exists>x. Q x)"
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   148
  by fast+}
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   149
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   150
fun quant_intro ctxt [thm] t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   151
    SMT_Replay_Methods.prove ctxt t (K (REPEAT_ALL_NEW (resolve_tac ctxt (thm :: quant_intro_rules))))
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   152
  | quant_intro ctxt thms t = replay_rule_error ctxt Z3_Proof.Quant_Intro thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   153
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   154
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   155
(* distributivity of conjunctions and disjunctions *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   156
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   157
(* TODO: there are no tests with this proof rule *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   158
fun distrib ctxt _ t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   159
  SMT_Replay_Methods.prove_abstract' ctxt t prop_tac
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   160
    (SMT_Replay_Methods.abstract_prop (dest_prop t))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   161
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   162
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
(* elimination of conjunctions *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   164
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
fun and_elim ctxt [thm] t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   166
      SMT_Replay_Methods.prove_abstract ctxt [thm] t prop_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   167
        SMT_Replay_Methods.abstract_lit (dest_prop t) ##>>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   168
        SMT_Replay_Methods.abstract_conj (dest_thm thm) #>>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   169
        apfst single o swap)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   170
  | and_elim ctxt thms t = replay_rule_error ctxt Z3_Proof.And_Elim thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   171
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   172
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   173
(* elimination of negated disjunctions *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   174
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   175
fun not_or_elim ctxt [thm] t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   176
      SMT_Replay_Methods.prove_abstract ctxt [thm] t prop_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   177
        SMT_Replay_Methods.abstract_lit (dest_prop t) ##>>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   178
        SMT_Replay_Methods.abstract_not SMT_Replay_Methods.abstract_disj (dest_thm thm) #>>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   179
        apfst single o swap)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   180
  | not_or_elim ctxt thms t =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   181
      replay_rule_error ctxt Z3_Proof.Not_Or_Elim thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   182
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   183
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   184
(* rewriting *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   185
57144
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   186
local
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   187
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
   188
fun dest_all (Const (\<^const_name>\<open>HOL.All\<close>, _) $ Abs (_, T, t)) nctxt =
57144
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   189
      let
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   190
        val (n, nctxt') = Name.variant "" nctxt
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   191
        val f = Free (n, T)
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   192
        val t' = Term.subst_bound (f, t)
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   193
      in dest_all t' nctxt' |>> cons f end
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   194
  | dest_all t _ = ([], t)
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   195
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   196
fun dest_alls t =
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   197
  let
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   198
    val nctxt = Name.make_context (Term.add_free_names t [])
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   199
    val (lhs, rhs) = HOLogic.dest_eq (dest_prop t)
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   200
    val (ls, lhs') = dest_all lhs nctxt
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   201
    val (rs, rhs') = dest_all rhs nctxt
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   202
  in
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   203
    if eq_list (op aconv) (ls, rs) then SOME (ls, (HOLogic.mk_eq (lhs', rhs')))
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   204
    else NONE
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   205
  end
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   206
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   207
fun forall_intr ctxt t thm =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
   208
  let val ct = Thm.cterm_of ctxt t
57144
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   209
  in Thm.forall_intr ct thm COMP_INCR @{thm iff_allI} end
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   210
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   211
in
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   212
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   213
fun focus_eq f ctxt t =
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   214
  (case dest_alls t of
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   215
    NONE => f ctxt t
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   216
  | SOME (vs, t') => fold (forall_intr ctxt) vs (f ctxt t'))
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   217
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   218
end
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   219
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
   220
fun abstract_eq f (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t1 $ t2) =
57144
1d12e22e7caf more complete proof replay for Z3: support universally quantified rewrite steps
boehmes
parents: 56816
diff changeset
   221
      f t1 ##>> f t2 #>> HOLogic.mk_eq
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   222
  | abstract_eq _ t = SMT_Replay_Methods.abstract_term t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   223
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   224
fun prove_prop_rewrite ctxt t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   225
  SMT_Replay_Methods.prove_abstract' ctxt t prop_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   226
    abstract_eq SMT_Replay_Methods.abstract_prop (dest_prop t))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   227
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   228
fun arith_rewrite_tac ctxt _ =
66692
00b54799bd29 strengthened reconstruction tactic
blanchet
parents: 65801
diff changeset
   229
  let val backup_tac = Arith_Data.arith_tac ctxt ORELSE' Clasimp.force_tac ctxt in
00b54799bd29 strengthened reconstruction tactic
blanchet
parents: 65801
diff changeset
   230
    (TRY o Simplifier.simp_tac ctxt) THEN_ALL_NEW backup_tac
00b54799bd29 strengthened reconstruction tactic
blanchet
parents: 65801
diff changeset
   231
    ORELSE' backup_tac
00b54799bd29 strengthened reconstruction tactic
blanchet
parents: 65801
diff changeset
   232
  end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   233
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   234
fun prove_arith_rewrite ctxt t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   235
  SMT_Replay_Methods.prove_abstract' ctxt t arith_rewrite_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   236
    abstract_eq (SMT_Replay_Methods.abstract_arith ctxt) (dest_prop t))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   237
58140
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   238
val lift_ite_thm = @{thm HOL.if_distrib} RS @{thm eq_reflection}
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   239
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   240
fun ternary_conv cv = Conv.combination_conv (Conv.binop_conv cv) cv
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   241
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   242
fun if_context_conv ctxt ct =
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   243
  (case Thm.term_of ct of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
   244
    Const (\<^const_name>\<open>HOL.If\<close>, _) $ _ $ _ $ _ =>
58140
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   245
      ternary_conv (if_context_conv ctxt)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
   246
  | _ $ (Const (\<^const_name>\<open>HOL.If\<close>, _) $ _ $ _ $ _) =>
58140
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   247
      Conv.rewr_conv lift_ite_thm then_conv ternary_conv (if_context_conv ctxt)
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   248
  | _ => Conv.sub_conv (Conv.top_sweep_conv if_context_conv) ctxt) ct
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   249
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   250
fun lift_ite_rewrite ctxt t =
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
   251
  SMT_Replay_Methods.prove ctxt t (fn ctxt =>
58140
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   252
    CONVERSION (HOLogic.Trueprop_conv (Conv.binop_conv (if_context_conv ctxt)))
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   253
    THEN' resolve_tac ctxt @{thms refl})
58140
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   254
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   255
fun prove_conj_disj_perm ctxt t = SMT_Replay_Methods.prove ctxt t Conj_Disj_Perm.conj_disj_perm_tac
59381
de4218223e00 more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
boehmes
parents: 58957
diff changeset
   256
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
   257
fun rewrite ctxt _ = try_provers ctxt
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   258
    (Z3_Proof.string_of_rule Z3_Proof.Rewrite) [
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   259
  ("rules", apply_rule ctxt),
59381
de4218223e00 more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
boehmes
parents: 58957
diff changeset
   260
  ("conj_disj_perm", prove_conj_disj_perm ctxt),
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   261
  ("prop_rewrite", prove_prop_rewrite ctxt),
58140
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   262
  ("arith_rewrite", focus_eq prove_arith_rewrite ctxt),
b4aa77aef6a8 replay Z3 rewrite steps that lift if-then-else expressions
boehmes
parents: 58061
diff changeset
   263
  ("if_rewrite", lift_ite_rewrite ctxt)] []
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   264
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   265
fun rewrite_star ctxt = rewrite ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   266
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   267
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   268
(* pulling quantifiers *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   269
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   270
fun pull_quant ctxt _ t = SMT_Replay_Methods.prove ctxt t quant_tac
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   271
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   272
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   273
(* pushing quantifiers *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   274
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   275
fun push_quant _ _ _ = raise Fail "unsupported" (* FIXME *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   276
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   277
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   278
(* elimination of unused bound variables *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   279
67091
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   280
val elim_all = @{lemma "P = Q \<Longrightarrow> (\<forall>x. P) = Q" by fast}
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   281
val elim_ex = @{lemma "P = Q \<Longrightarrow> (\<exists>x. P) = Q" by fast}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   282
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58956
diff changeset
   283
fun elim_unused_tac ctxt i st = (
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58956
diff changeset
   284
  match_tac ctxt [@{thm refl}]
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58956
diff changeset
   285
  ORELSE' (match_tac ctxt [elim_all, elim_ex] THEN' elim_unused_tac ctxt)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   286
  ORELSE' (
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58956
diff changeset
   287
    match_tac ctxt [@{thm iff_allI}, @{thm iff_exI}]
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58956
diff changeset
   288
    THEN' elim_unused_tac ctxt)) i st
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   289
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   290
fun elim_unused ctxt _ t = SMT_Replay_Methods.prove ctxt t elim_unused_tac
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   291
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   292
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   293
(* destructive equality resolution *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   294
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   295
fun dest_eq_res _ _ _ = raise Fail "dest_eq_res" (* FIXME *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   296
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   297
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   298
(* quantifier instantiation *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   299
67091
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   300
val quant_inst_rule = @{lemma "\<not>P x \<or> Q ==> \<not>(\<forall>x. P x) \<or> Q" by fast}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   301
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   302
fun quant_inst ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ =>
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   303
  REPEAT_ALL_NEW (resolve_tac ctxt [quant_inst_rule])
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   304
  THEN' resolve_tac ctxt @{thms excluded_middle})
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   305
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   306
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   307
(* propositional lemma *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   308
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   309
exception LEMMA of unit
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   310
67091
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   311
val intro_hyp_rule1 = @{lemma "(\<not>P \<Longrightarrow> Q) \<Longrightarrow> P \<or> Q" by fast}
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   312
val intro_hyp_rule2 = @{lemma "(P \<Longrightarrow> Q) \<Longrightarrow> \<not>P \<or> Q" by fast}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   313
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   314
fun norm_lemma thm =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   315
  (thm COMP_INCR intro_hyp_rule1)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   316
  handle THM _ => thm COMP_INCR intro_hyp_rule2
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   317
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 72458
diff changeset
   318
fun negated_prop \<^Const_>\<open>Not for t\<close> = HOLogic.mk_Trueprop t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   319
  | negated_prop t = HOLogic.mk_Trueprop (HOLogic.mk_not t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   320
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 72458
diff changeset
   321
fun intro_hyps tab (t as \<^Const_>\<open>HOL.disj\<close> $ t1 $ t2) cx =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   322
      lookup_intro_hyps tab t (fold (intro_hyps tab) [t1, t2]) cx
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   323
  | intro_hyps tab t cx =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   324
      lookup_intro_hyps tab t (fn _ => raise LEMMA ()) cx
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   325
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   326
and lookup_intro_hyps tab t f (cx as (thm, terms)) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   327
  (case Termtab.lookup tab (negated_prop t) of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   328
    NONE => f cx
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   329
  | SOME hyp => (norm_lemma (Thm.implies_intr hyp thm), t :: terms))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   330
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   331
fun lemma ctxt (thms as [thm]) t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   332
    (let
60949
ccbf9379e355 added Thm.chyps_of;
wenzelm
parents: 60752
diff changeset
   333
       val tab = Termtab.make (map (`Thm.term_of) (Thm.chyps_of thm))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   334
       val (thm', terms) = intro_hyps tab (dest_prop t) (thm, [])
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   335
     in
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   336
       SMT_Replay_Methods.prove_abstract ctxt [thm'] t prop_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   337
         fold (snd oo SMT_Replay_Methods.abstract_lit) terms #>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   338
         SMT_Replay_Methods.abstract_disj (dest_thm thm') #>> single ##>>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   339
         SMT_Replay_Methods.abstract_disj (dest_prop t))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   340
     end
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   341
     handle LEMMA () => replay_error ctxt "Bad proof state" Z3_Proof.Lemma thms t)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   342
  | lemma ctxt thms t = replay_rule_error ctxt Z3_Proof.Lemma thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   343
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   344
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   345
(* unit resolution *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   346
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   347
fun unit_res ctxt thms t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   348
  SMT_Replay_Methods.prove_abstract ctxt thms t prop_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   349
    fold_map (SMT_Replay_Methods.abstract_unit o dest_thm) thms ##>>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   350
    SMT_Replay_Methods.abstract_unit (dest_prop t) #>>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   351
    (fn (prems, concl) => (prems, concl)))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   352
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   353
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   354
(* iff-true *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   355
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   356
val iff_true_rule = @{lemma "P ==> P = True" by fast}
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   357
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   358
fun iff_true _ [thm] _ = thm RS iff_true_rule
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   359
  | iff_true ctxt thms t = replay_rule_error ctxt Z3_Proof.Iff_True thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   360
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   361
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   362
(* iff-false *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   363
67091
1393c2340eec more symbols;
wenzelm
parents: 66692
diff changeset
   364
val iff_false_rule = @{lemma "\<not>P \<Longrightarrow> P = False" by fast}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   365
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   366
fun iff_false _ [thm] _ = thm RS iff_false_rule
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   367
  | iff_false ctxt thms t = replay_rule_error ctxt Z3_Proof.Iff_False thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   368
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   369
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   370
(* commutativity *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   371
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   372
fun comm ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t @{thm eq_commute}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   373
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   374
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   375
(* definitional axioms *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   376
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   377
fun def_axiom_disj ctxt t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   378
  (case dest_prop t of
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 72458
diff changeset
   379
    \<^Const_>\<open>disj for u1 u2\<close> =>
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   380
      SMT_Replay_Methods.prove_abstract' ctxt t prop_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   381
         SMT_Replay_Methods.abstract_prop u2 ##>>  SMT_Replay_Methods.abstract_prop u1 #>>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   382
         HOLogic.mk_disj o swap)
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   383
  | u => SMT_Replay_Methods.prove_abstract' ctxt t prop_tac (SMT_Replay_Methods.abstract_prop u))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   384
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
   385
fun def_axiom ctxt _ = try_provers ctxt
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   386
    (Z3_Proof.string_of_rule Z3_Proof.Def_Axiom) [
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   387
  ("rules", apply_rule ctxt),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   388
  ("disj", def_axiom_disj ctxt)] []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   389
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   390
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   391
(* application of definitions *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   392
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   393
fun apply_def _ [thm] _ = thm (* TODO: cover also the missing cases *)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   394
  | apply_def ctxt thms t = replay_rule_error ctxt Z3_Proof.Apply_Def thms t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   395
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   396
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   397
(* iff-oeq *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   398
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   399
fun iff_oeq _ _ _ = raise Fail "iff_oeq" (* FIXME *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   400
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   401
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   402
(* negation normal form *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   403
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   404
fun nnf_prop ctxt thms t =
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   405
  SMT_Replay_Methods.prove_abstract ctxt thms t prop_tac (
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   406
    fold_map (SMT_Replay_Methods.abstract_prop o dest_thm) thms ##>>
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67091
diff changeset
   407
    SMT_Replay_Methods.abstract_prop (dest_prop t))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   408
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 69597
diff changeset
   409
fun nnf ctxt rule thms = try_provers ctxt (Z3_Proof.string_of_rule rule) [
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   410
  ("prop", nnf_prop ctxt thms),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   411
  ("quant", quant_intro ctxt [hd thms])] thms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   412
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   413
fun nnf_pos ctxt = nnf ctxt Z3_Proof.Nnf_Pos
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   414
fun nnf_neg ctxt = nnf ctxt Z3_Proof.Nnf_Neg
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   415
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   416
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   417
(* mapping of rules to methods *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   418
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   419
fun unsupported rule ctxt = replay_error ctxt "Unsupported" rule
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   420
fun assumed rule ctxt = replay_error ctxt "Assumed" rule
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   421
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   422
fun choose Z3_Proof.True_Axiom = true_axiom
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   423
  | choose (r as Z3_Proof.Asserted) = assumed r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   424
  | choose (r as Z3_Proof.Goal) = assumed r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   425
  | choose Z3_Proof.Modus_Ponens = mp
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   426
  | choose Z3_Proof.Reflexivity = refl
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   427
  | choose Z3_Proof.Symmetry = symm
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   428
  | choose Z3_Proof.Transitivity = trans
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   429
  | choose (r as Z3_Proof.Transitivity_Star) = unsupported r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   430
  | choose Z3_Proof.Monotonicity = cong
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   431
  | choose Z3_Proof.Quant_Intro = quant_intro
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   432
  | choose Z3_Proof.Distributivity = distrib
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   433
  | choose Z3_Proof.And_Elim = and_elim
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   434
  | choose Z3_Proof.Not_Or_Elim = not_or_elim
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   435
  | choose Z3_Proof.Rewrite = rewrite
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   436
  | choose Z3_Proof.Rewrite_Star = rewrite_star
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   437
  | choose Z3_Proof.Pull_Quant = pull_quant
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   438
  | choose (r as Z3_Proof.Pull_Quant_Star) = unsupported r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   439
  | choose Z3_Proof.Push_Quant = push_quant
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   440
  | choose Z3_Proof.Elim_Unused_Vars = elim_unused
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   441
  | choose Z3_Proof.Dest_Eq_Res = dest_eq_res
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   442
  | choose Z3_Proof.Quant_Inst = quant_inst
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   443
  | choose (r as Z3_Proof.Hypothesis) = assumed r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   444
  | choose Z3_Proof.Lemma = lemma
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   445
  | choose Z3_Proof.Unit_Resolution = unit_res
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   446
  | choose Z3_Proof.Iff_True = iff_true
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   447
  | choose Z3_Proof.Iff_False = iff_false
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   448
  | choose Z3_Proof.Commutativity = comm
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   449
  | choose Z3_Proof.Def_Axiom = def_axiom
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   450
  | choose (r as Z3_Proof.Intro_Def) = assumed r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   451
  | choose Z3_Proof.Apply_Def = apply_def
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   452
  | choose Z3_Proof.Iff_Oeq = iff_oeq
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   453
  | choose Z3_Proof.Nnf_Pos = nnf_pos
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   454
  | choose Z3_Proof.Nnf_Neg = nnf_neg
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   455
  | choose (r as Z3_Proof.Nnf_Star) = unsupported r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   456
  | choose (r as Z3_Proof.Cnf_Star) = unsupported r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   457
  | choose (r as Z3_Proof.Skolemize) = assumed r
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   458
  | choose Z3_Proof.Modus_Ponens_Oeq = mp_oeq
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   459
  | choose (Z3_Proof.Th_Lemma name) = th_lemma name
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   460
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   461
fun with_tracing rule method ctxt thms t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   462
  let val _ = trace_goal ctxt rule thms t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   463
  in method ctxt thms t end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   464
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   465
fun method_for rule = with_tracing rule (choose rule)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   466
57229
blanchet
parents: 57220
diff changeset
   467
end;