| author | wenzelm | 
| Mon, 21 Mar 2016 20:22:07 +0100 | |
| changeset 62681 | 45b8dd2d3827 | 
| parent 62519 | a564458f94db | 
| child 67091 | 1393c2340eec | 
| permissions | -rw-r--r-- | 
| 58061 | 1 | (* Title: HOL/Tools/SMT/z3_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: 
57165diff
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 | |
| 58061 | 8 | signature Z3_REPLAY = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 9 | sig | 
| 58061 | 10 | val parse_proof: Proof.context -> SMT_Translate.replay_data -> | 
| 57164 | 11 | ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list -> | 
| 58061 | 12 | SMT_Solver.parsed_proof | 
| 13 | val replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm | |
| 57229 | 14 | end; | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 15 | |
| 58061 | 16 | structure Z3_Replay: Z3_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 | 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 | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 26 | in Drule.forall_elim_list (map (Thm.cterm_of 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 | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 34 | fun mk (n, T) n' = (n, Thm.cterm_of 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 | 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 | |
| 58061 | 59 | fun replay_thm ctxt assumed nthms (Z3_Proof.Z3_Step {id, rule, concl, fixes, is_fix_step, ...}) =
 | 
| 60 | if Z3_Proof.is_assumption rule then | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 61 | (case Inttab.lookup assumed id of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 | SOME (_, thm) => thm | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 63 | | NONE => Thm.assume (Thm.cterm_of ctxt concl)) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 64 | else | 
| 58061 | 65 | under_fixes (Z3_Replay_Methods.method_for rule) ctxt | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 66 | (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 | 67 | |
| 59215 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 68 | fun replay_step ctxt assumed (step as Z3_Proof.Z3_Step {id, rule, prems, fixes, ...}) state =
 | 
| 59213 
ef5e68575bc4
limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
 boehmes parents: 
58482diff
changeset | 69 | let | 
| 59215 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 70 | val (proofs, stats) = state | 
| 59213 
ef5e68575bc4
limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
 boehmes parents: 
58482diff
changeset | 71 | val nthms = map (the o Inttab.lookup proofs) prems | 
| 59215 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 72 | val replay = Timing.timing (replay_thm ctxt assumed nthms) | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 73 |     val ({elapsed, ...}, thm) =
 | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 74 | SMT_Config.with_time_limit ctxt SMT_Config.reconstruction_step_timeout replay step | 
| 62519 | 75 | handle Timeout.TIMEOUT _ => raise SMT_Failure.SMT SMT_Failure.Time_Out | 
| 59215 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 76 | val stats' = Symtab.cons_list (Z3_Proof.string_of_rule rule, Time.toMilliseconds elapsed) stats | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 77 | in (Inttab.update (id, (fixes, thm)) proofs, stats') end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 78 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 79 | local | 
| 57230 | 80 |   val remove_trigger = mk_meta_eq @{thm trigger_def}
 | 
| 81 |   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 | 82 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 83 | fun rewrite_conv _ [] = Conv.all_conv | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 | | 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 | 85 | |
| 59381 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 boehmes parents: 
59374diff
changeset | 86 |   val rewrite_true_rule = @{lemma "True == ~ False" by simp}
 | 
| 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 boehmes parents: 
59374diff
changeset | 87 |   val prep_rules = [@{thm Let_def}, remove_trigger, remove_fun_app, rewrite_true_rule]
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 88 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 89 | fun rewrite _ [] = I | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 90 | | 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 | 91 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 92 | fun lookup_assm assms_net ct = | 
| 58061 | 93 | Z3_Replay_Util.net_instances assms_net ct | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 94 | |> 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 | 95 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 96 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 97 | 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 | 98 | let | 
| 59381 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 boehmes parents: 
59374diff
changeset | 99 | val eqs = map (rewrite ctxt [rewrite_true_rule]) rewrite_rules | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 100 | val eqs' = union Thm.eq_thm eqs prep_rules | 
| 
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 | val assms_net = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 103 | assms | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 104 | |> map (apsnd (rewrite ctxt eqs')) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 105 | |> map (apsnd (Conv.fconv_rule Thm.eta_conversion)) | 
| 58061 | 106 | |> Z3_Replay_Util.thm_net_of snd | 
| 56078 
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 | 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 | 109 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 110 | fun assume thm ctxt = | 
| 
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 ct = Thm.cprem_of thm 1 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 113 | 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 | 114 | in (thm' RS thm, ctxt') end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 115 | |
| 56104 
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
 blanchet parents: 
56099diff
changeset | 116 | 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 | 117 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 118 | 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 | 119 | val thms' = if exact then thms else th :: thms | 
| 56104 
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
 blanchet parents: 
56099diff
changeset | 120 | 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 | 121 | |
| 58061 | 122 |     fun add (Z3_Proof.Z3_Step {id, rule, concl, fixes, ...})
 | 
| 56104 
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
 blanchet parents: 
56099diff
changeset | 123 | (cx as ((iidths, thms), (ctxt, ptab))) = | 
| 58061 | 124 | if Z3_Proof.is_assumption rule andalso rule <> Z3_Proof.Hypothesis then | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 125 | let | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 126 | val ct = Thm.cterm_of ctxt concl | 
| 56099 | 127 | 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 | 128 | val thm2 = singleton (Variable.export ctxt outer_ctxt) thm1 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 129 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 130 | (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 | 131 | [] => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 132 | let val (thm, ctxt') = assume thm1 ctxt | 
| 56104 
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
 blanchet parents: 
56099diff
changeset | 133 | 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 | 134 | | ithms => fold (add1 id fixes thm1) ithms cx) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 135 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 136 | else | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 137 | cx | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 138 | in fold add steps (([], []), (ctxt, Inttab.empty)) end | 
| 
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 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 141 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 142 | (* |- (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 | 143 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 144 |   val sk_rules = @{lemma
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 145 | "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 | 146 | "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 | 147 | by (metis someI_ex)+} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 148 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 149 | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59381diff
changeset | 150 | fun discharge_sk_tac ctxt i st = | 
| 60752 | 151 |   (resolve_tac ctxt @{thms trans} i
 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59381diff
changeset | 152 | THEN resolve_tac ctxt sk_rules i | 
| 60752 | 153 |    THEN (resolve_tac ctxt @{thms refl} ORELSE' discharge_sk_tac ctxt) (i+1)
 | 
| 154 |    THEN resolve_tac ctxt @{thms refl} i) st
 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 155 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 156 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 157 | |
| 59381 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 boehmes parents: 
59374diff
changeset | 158 | val true_thm = @{lemma "~False" by simp}
 | 
| 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 boehmes parents: 
59374diff
changeset | 159 | fun make_discharge_rules rules = rules @ [@{thm allI}, @{thm refl}, @{thm reflexive}, true_thm]
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 160 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 161 | val intro_def_rules = @{lemma
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 162 | "(~ P | P) & (P | ~ P)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 163 | "(P | ~ P) & (~ P | P)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 164 | by fast+} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 165 | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59381diff
changeset | 166 | fun discharge_assms_tac ctxt rules = | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59381diff
changeset | 167 | REPEAT | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59381diff
changeset | 168 | (HEADGOAL (resolve_tac ctxt (intro_def_rules @ rules) ORELSE' | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59381diff
changeset | 169 | SOLVED' (discharge_sk_tac ctxt))) | 
| 57230 | 170 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 171 | fun discharge_assms ctxt rules thm = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 172 | (if Thm.nprems_of thm = 0 then | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 173 | thm | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 174 | else | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59381diff
changeset | 175 | (case Seq.pull (discharge_assms_tac ctxt rules thm) of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 176 | SOME (thm', _) => thm' | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 177 |      | NONE => raise THM ("failed to discharge premise", 1, [thm])))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 178 | |> Goal.norm_result ctxt | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 179 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 180 | fun discharge rules outer_ctxt inner_ctxt = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 181 | singleton (Proof_Context.export inner_ctxt outer_ctxt) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 182 | #> discharge_assms outer_ctxt (make_discharge_rules rules) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 183 | |
| 57157 | 184 | fun parse_proof outer_ctxt | 
| 58061 | 185 |     ({context = ctxt, typs, terms, ll_defs, rewrite_rules, assms} : SMT_Translate.replay_data)
 | 
| 57541 | 186 | xfacts prems concl output = | 
| 57157 | 187 | let | 
| 58061 | 188 | val (steps, ctxt2) = Z3_Proof.parse typs terms output ctxt | 
| 57157 | 189 | val ((iidths, _), _) = add_asserted outer_ctxt rewrite_rules assms steps ctxt2 | 
| 57164 | 190 | |
| 191 | fun id_of_index i = the_default ~1 (Option.map fst (AList.lookup (op =) iidths i)) | |
| 192 | ||
| 193 | val conjecture_i = 0 | |
| 194 | val prems_i = 1 | |
| 195 | val facts_i = prems_i + length prems | |
| 196 | ||
| 197 | val conjecture_id = id_of_index conjecture_i | |
| 198 | val prem_ids = map id_of_index (prems_i upto facts_i - 1) | |
| 58482 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58061diff
changeset | 199 | val fact_ids' = | 
| 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58061diff
changeset | 200 | map_filter (fn (i, (id, _)) => try (apsnd (nth xfacts)) (id, i - facts_i)) iidths | 
| 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58061diff
changeset | 201 | val helper_ids' = map_filter (try (fn (~1, idth) => idth)) iidths | 
| 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58061diff
changeset | 202 | |
| 57164 | 203 | val fact_helper_ts = | 
| 59582 | 204 | map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, Thm.prop_of th)) helper_ids' @ | 
| 205 | map (fn (_, ((s, _), th)) => (s, Thm.prop_of th)) fact_ids' | |
| 58482 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58061diff
changeset | 206 | val fact_helper_ids' = | 
| 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58061diff
changeset | 207 | map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids' @ map (apsnd (fst o fst)) fact_ids' | 
| 57157 | 208 | in | 
| 60201 | 209 |     {outcome = NONE, fact_ids = SOME fact_ids',
 | 
| 58061 | 210 | atp_proof = fn () => Z3_Isar.atp_proof_of_z3_proof ctxt ll_defs rewrite_rules prems concl | 
| 58482 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58061diff
changeset | 211 | fact_helper_ts prem_ids conjecture_id fact_helper_ids' steps} | 
| 57157 | 212 | end | 
| 213 | ||
| 59374 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 214 | fun intermediate_statistics ctxt start total = | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 215 | SMT_Config.statistics_msg ctxt (fn current => | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 216 | "Reconstructed " ^ string_of_int current ^ " of " ^ string_of_int total ^ " steps in " ^ | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 217 | string_of_int (Time.toMilliseconds (#elapsed (Timing.result start))) ^ " ms") | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 218 | |
| 59215 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 219 | fun pretty_statistics total stats = | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 220 | let | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 221 | fun mean_of is = | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 222 | let | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 223 | val len = length is | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 224 | val mid = len div 2 | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 225 | in if len mod 2 = 0 then (nth is (mid - 1) + nth is mid) div 2 else nth is mid end | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 226 | fun pretty_item name p = Pretty.item (Pretty.separate ":" [Pretty.str name, p]) | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 227 | fun pretty (name, milliseconds) = pretty_item name (Pretty.block (Pretty.separate "," [ | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 228 | Pretty.str (string_of_int (length milliseconds) ^ " occurrences") , | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 229 | Pretty.str (string_of_int (mean_of milliseconds) ^ " ms mean time"), | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 230 | Pretty.str (string_of_int (fold Integer.max milliseconds 0) ^ " ms maximum time"), | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 231 | Pretty.str (string_of_int (fold Integer.add milliseconds 0) ^ " ms total time")])) | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 232 | in | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 233 | Pretty.big_list "Z3 proof reconstruction statistics:" ( | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 234 | pretty_item "total time" (Pretty.str (string_of_int total ^ " ms")) :: | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 235 | map pretty (Symtab.dest stats)) | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 236 | end | 
| 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 237 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 238 | fun replay outer_ctxt | 
| 58061 | 239 |     ({context = ctxt, typs, terms, rewrite_rules, assms, ...} : SMT_Translate.replay_data) output =
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 240 | let | 
| 58061 | 241 | val (steps, ctxt2) = Z3_Proof.parse typs terms output ctxt | 
| 57157 | 242 | val ((_, rules), (ctxt3, assumed)) = add_asserted outer_ctxt rewrite_rules assms steps ctxt2 | 
| 243 | val ctxt4 = | |
| 244 | ctxt3 | |
| 58061 | 245 | |> put_simpset (Z3_Replay_Util.make_simpset ctxt3 []) | 
| 246 | |> Config.put SAT.solver (Config.get ctxt3 SMT_Config.sat_solver) | |
| 59374 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 247 | val len = length steps | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 248 | val start = Timing.start () | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 249 | val print_runtime_statistics = intermediate_statistics ctxt4 start len | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 250 | fun blockwise f (i, x) y = | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 251 | (if i > 0 andalso i mod 100 = 0 then print_runtime_statistics i else (); f x y) | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 252 | val (proofs, stats) = | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 253 | fold_index (blockwise (replay_step ctxt4 assumed)) steps (assumed, Symtab.empty) | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 254 | val _ = print_runtime_statistics len | 
| 
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
 boehmes parents: 
59215diff
changeset | 255 | val total = Time.toMilliseconds (#elapsed (Timing.result start)) | 
| 58061 | 256 |     val (_, Z3_Proof.Z3_Step {id, ...}) = split_last steps
 | 
| 59215 
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
 boehmes parents: 
59213diff
changeset | 257 | val _ = SMT_Config.statistics_msg ctxt4 (Pretty.string_of o pretty_statistics total) stats | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 258 | in | 
| 57157 | 259 | 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 | 260 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 261 | |
| 57229 | 262 | end; |