src/HOL/Tools/SMT2/z3_new_replay.ML
author blanchet
Thu, 12 Jun 2014 01:00:49 +0200
changeset 57230 ec5ff6bb2a92
parent 57229 489083abce44
child 57541 147e3f1e0459
permissions -rw-r--r--
eliminate dependency of SMT2 module on 'list'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
     1
(*  Title:      HOL/Tools/SMT2/z3_new_replay.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
57219
34018603e0d0 factor out SMT-LIB 2 type/term parsing from Z3-specific code
blanchet
parents: 57165
diff changeset
     5
Z3 proof parsing and replay.
56078
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
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
     8
signature Z3_NEW_REPLAY =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     9
sig
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
    10
  val parse_proof: Proof.context -> SMT2_Translate.replay_data ->
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
    11
    ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
    12
    SMT2_Solver.parsed_proof
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
    13
  val replay: Proof.context -> SMT2_Translate.replay_data -> string list -> thm
57229
blanchet
parents: 57220
diff changeset
    14
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
    16
structure Z3_New_Replay: Z3_NEW_REPLAY =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    18
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 56104
diff changeset
    19
fun params_of t = Term.strip_qnt_vars @{const_name Pure.all} t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    20
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    21
fun varify ctxt thm =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    22
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    23
    val maxidx = Thm.maxidx_of thm + 1
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    24
    val vs = params_of (Thm.prop_of thm)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
    val vars = map_index (fn (i, (n, T)) => Var ((n, i + maxidx), T)) vs
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
    26
  in Drule.forall_elim_list (map (SMT2_Util.certify ctxt) vars) thm end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    27
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
fun add_paramTs names t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    29
  fold2 (fn n => fn (_, T) => AList.update (op =) (n, T)) names (params_of t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    30
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
fun new_fixes ctxt nTs =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    32
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
    val (ns, ctxt') = Variable.variant_fixes (replicate (length nTs) "") ctxt
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
    34
    fun mk (n, T) n' = (n, SMT2_Util.certify ctxt' (Free (n', T)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
  in (ctxt', Symtab.make (map2 mk nTs ns)) end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    36
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 56104
diff changeset
    37
fun forall_elim_term ct (Const (@{const_name Pure.all}, _) $ (a as Abs _)) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    38
      Term.betapply (a, Thm.term_of ct)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    39
  | forall_elim_term _ qt = raise TERM ("forall_elim'", [qt])
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    41
fun apply_fixes elim env = fold (elim o the o Symtab.lookup env)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    42
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
val apply_fixes_prem = uncurry o apply_fixes Thm.forall_elim
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    44
val apply_fixes_concl = apply_fixes forall_elim_term
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    45
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    46
fun export_fixes env names = Drule.forall_intr_list (map (the o Symtab.lookup env) names)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    47
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    48
fun under_fixes f ctxt (prems, nthms) names concl =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    49
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
    val thms1 = map (varify ctxt) prems
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
    val (ctxt', env) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
      add_paramTs names concl []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
      |> fold (uncurry add_paramTs o apsnd Thm.prop_of) nthms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    54
      |> new_fixes ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    55
    val thms2 = map (apply_fixes_prem env) nthms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    56
    val t = apply_fixes_concl env names concl
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
  in export_fixes env names (f ctxt' (thms1 @ thms2) t) end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
fun replay_thm ctxt assumed nthms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    60
    (Z3_New_Proof.Z3_Step {id, rule, concl, fixes, is_fix_step, ...}) =
57220
853557cf2efa tuned code
blanchet
parents: 57219
diff changeset
    61
  if Z3_New_Proof.is_assumption rule then
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
    (case Inttab.lookup assumed id of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    63
      SOME (_, thm) => thm
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
    64
    | NONE => Thm.assume (SMT2_Util.certify ctxt concl))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
  else
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
    66
    under_fixes (Z3_New_Replay_Methods.method_for rule) ctxt
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    67
      (if is_fix_step then (map snd nthms, []) else ([], nthms)) fixes concl
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    68
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    69
fun replay_step ctxt assumed (step as Z3_New_Proof.Z3_Step {id, prems, fixes, ...}) proofs =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    70
  let val nthms = map (the o Inttab.lookup proofs) prems
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    71
  in Inttab.update (id, (fixes, replay_thm ctxt assumed nthms step)) proofs end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    72
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
local
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
    74
  val remove_trigger = mk_meta_eq @{thm trigger_def}
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
    75
  val remove_fun_app = mk_meta_eq @{thm fun_app_def}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    76
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    77
  fun rewrite_conv _ [] = Conv.all_conv
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    78
    | rewrite_conv ctxt eqs = Simplifier.full_rewrite (empty_simpset ctxt addsimps eqs)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    79
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57164
diff changeset
    80
  val prep_rules = [@{thm Let_def}, remove_trigger, remove_fun_app,
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57164
diff changeset
    81
    Z3_New_Replay_Literals.rewrite_true]
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    82
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    83
  fun rewrite _ [] = I
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    84
    | rewrite ctxt eqs = Conv.fconv_rule (rewrite_conv ctxt eqs)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    85
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    86
  fun lookup_assm assms_net ct =
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
    87
    Z3_New_Replay_Util.net_instances assms_net ct
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
    |> map (fn ithm as (_, thm) => (ithm, Thm.cprop_of thm aconvc ct))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    89
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    90
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    91
fun add_asserted outer_ctxt rewrite_rules assms steps ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    92
  let
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
    93
    val eqs = map (rewrite ctxt [Z3_New_Replay_Literals.rewrite_true]) rewrite_rules
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    94
    val eqs' = union Thm.eq_thm eqs prep_rules
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    95
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    96
    val assms_net =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    97
      assms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    98
      |> map (apsnd (rewrite ctxt eqs'))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    99
      |> map (apsnd (Conv.fconv_rule Thm.eta_conversion))
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   100
      |> Z3_New_Replay_Util.thm_net_of snd
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
    fun revert_conv ctxt = rewrite_conv ctxt eqs' then_conv Thm.eta_conversion
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   104
    fun assume thm ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
      let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
        val ct = Thm.cprem_of thm 1
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
        val (thm', ctxt') = yield_singleton Assumption.add_assumes ct ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   108
      in (thm' RS thm, ctxt') end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   109
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   110
    fun add1 id fixes thm1 ((i, th), exact) ((iidths, thms), (ctxt, ptab)) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   111
      let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   112
        val (thm, ctxt') = if exact then (Thm.implies_elim thm1 th, ctxt) else assume thm1 ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   113
        val thms' = if exact then thms else th :: thms
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   114
      in (((i, (id, th)) :: iidths, thms'), (ctxt', Inttab.update (id, (fixes, thm)) ptab)) end
56078
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
    fun add (Z3_New_Proof.Z3_Step {id, rule, concl, fixes, ...})
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   117
        (cx as ((iidths, thms), (ctxt, ptab))) =
57220
853557cf2efa tuned code
blanchet
parents: 57219
diff changeset
   118
      if Z3_New_Proof.is_assumption rule andalso rule <> Z3_New_Proof.Hypothesis then
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   119
        let
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
   120
          val ct = SMT2_Util.certify ctxt concl
56099
bc036c1cf111 thread through step IDs from Z3 to Sledgehammer
blanchet
parents: 56096
diff changeset
   121
          val thm1 = Thm.trivial ct |> Conv.fconv_rule (Conv.arg1_conv (revert_conv outer_ctxt))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   122
          val thm2 = singleton (Variable.export ctxt outer_ctxt) thm1
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   123
        in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   124
          (case lookup_assm assms_net (Thm.cprem_of thm2 1) of
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
              let val (thm, ctxt') = assume thm1 ctxt
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   127
              in ((iidths, thms), (ctxt', Inttab.update (id, (fixes, thm)) ptab)) end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   128
          | ithms => fold (add1 id fixes thm1) ithms cx)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   129
        end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   130
      else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   131
        cx
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   132
  in fold add steps (([], []), (ctxt, Inttab.empty)) end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   133
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   134
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   135
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   136
(* |- (EX x. P x) = P c     |- ~ (ALL x. P x) = ~ P c *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   137
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   138
  val sk_rules = @{lemma
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   139
    "c = (SOME x. P x) ==> (EX x. P x) = P c"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   140
    "c = (SOME x. ~ P x) ==> (~ (ALL x. P x)) = (~ P c)"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
    by (metis someI_ex)+}
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   142
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   143
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   144
fun discharge_sk_tac i st =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   145
  (rtac @{thm trans} i
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   146
   THEN resolve_tac sk_rules i
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   147
   THEN (rtac @{thm refl} ORELSE' discharge_sk_tac) (i+1)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   148
   THEN rtac @{thm refl} i) st
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
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   151
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   152
fun make_discharge_rules rules = rules @ [@{thm allI}, @{thm refl},
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56089
diff changeset
   153
  @{thm reflexive}, Z3_New_Replay_Literals.true_thm]
56078
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
val intro_def_rules = @{lemma
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   156
  "(~ P | P) & (P | ~ P)"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   157
  "(P | ~ P) & (~ P | P)"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   158
  by fast+}
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   159
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   160
fun discharge_assms_tac rules =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   161
  REPEAT (HEADGOAL (resolve_tac (intro_def_rules @ rules) ORELSE' SOLVED' discharge_sk_tac))
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   162
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
fun discharge_assms ctxt rules thm =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   164
  (if Thm.nprems_of thm = 0 then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
     thm
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
   else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   167
     (case Seq.pull (discharge_assms_tac rules thm) of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   168
       SOME (thm', _) => thm'
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   169
     | NONE => raise THM ("failed to discharge premise", 1, [thm])))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   170
  |> Goal.norm_result ctxt
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
fun discharge rules outer_ctxt inner_ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   173
  singleton (Proof_Context.export inner_ctxt outer_ctxt)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   174
  #> discharge_assms outer_ctxt (make_discharge_rules rules)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   175
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   176
fun parse_proof outer_ctxt
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   177
    ({context = ctxt, typs, terms, rewrite_rules, assms} : SMT2_Translate.replay_data) xfacts prems
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   178
    concl output =
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   179
  let
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   180
    val (steps, ctxt2) = Z3_New_Proof.parse typs terms output ctxt
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   181
    val ((iidths, _), _) = add_asserted outer_ctxt rewrite_rules assms steps ctxt2
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   182
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   183
    fun id_of_index i = the_default ~1 (Option.map fst (AList.lookup (op =) iidths i))
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   184
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   185
    val conjecture_i = 0
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   186
    val prems_i = 1
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   187
    val facts_i = prems_i + length prems
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   188
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   189
    val conjecture_id = id_of_index conjecture_i
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   190
    val prem_ids = map id_of_index (prems_i upto facts_i - 1)
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   191
    val helper_ids = map_filter (try (fn (~1, idth) => idth)) iidths
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   192
    val fact_ids = map_filter (fn (i, (id, _)) => try (apsnd (nth xfacts)) (id, i - facts_i)) iidths
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   193
    val fact_helper_ts =
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   194
      map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, prop_of th)) helper_ids @
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   195
      map (fn (_, ((s, _), th)) => (s, prop_of th)) fact_ids
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   196
    val fact_helper_ids =
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   197
      map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids @ map (apsnd (fst o fst)) fact_ids
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   198
  in
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   199
    {outcome = NONE, fact_ids = fact_ids,
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   200
     atp_proof = fn () => Z3_New_Isar.atp_proof_of_z3_proof ctxt rewrite_rules prems concl
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57157
diff changeset
   201
       fact_helper_ts prem_ids conjecture_id fact_helper_ids steps}
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   202
  end
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   203
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   204
fun replay outer_ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   205
    ({context=ctxt, typs, terms, rewrite_rules, assms} : SMT2_Translate.replay_data) output =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   206
  let
56095
f68150e2ead3 do less work in 'filter' mode
blanchet
parents: 56090
diff changeset
   207
    val (steps, ctxt2) = Z3_New_Proof.parse typs terms output ctxt
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   208
    val ((_, rules), (ctxt3, assumed)) = add_asserted outer_ctxt rewrite_rules assms steps ctxt2
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   209
    val ctxt4 =
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   210
      ctxt3
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   211
      |> put_simpset (Z3_New_Replay_Util.make_simpset ctxt3 [])
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   212
      |> Config.put SAT.solver (Config.get ctxt3 SMT2_Config.sat_solver)
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   213
    val proofs = fold (replay_step ctxt4 assumed) steps assumed
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   214
    val (_, Z3_New_Proof.Z3_Step {id, ...}) = split_last steps
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   215
  in
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 56816
diff changeset
   216
    Inttab.lookup proofs id |> the |> snd |> discharge rules outer_ctxt ctxt4
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   217
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   218
57229
blanchet
parents: 57220
diff changeset
   219
end;