| author | haftmann | 
| Sat, 28 Jun 2014 09:16:42 +0200 | |
| changeset 57418 | 6ab1c7cb0b8d | 
| parent 57289 | 5483868da0d8 | 
| child 57541 | 147e3f1e0459 | 
| permissions | -rw-r--r-- | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 1 | (* Title: HOL/Tools/SMT2/z3_new_isar.ML | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 2 | Author: Jasmin Blanchette, TU Muenchen | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 3 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 4 | Z3 proofs as generic ATP proofs for Isar proof reconstruction. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 5 | *) | 
| 
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 | signature Z3_NEW_ISAR = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 8 | sig | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 9 | val atp_proof_of_z3_proof: Proof.context -> thm list -> term list -> term -> | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 10 | (string * term) list -> int list -> int -> (int * string) list -> Z3_New_Proof.z3_step list -> | 
| 57159 | 11 | (term, string) ATP_Proof.atp_step list | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 12 | end; | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 13 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 14 | structure Z3_New_Isar: Z3_NEW_ISAR = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 15 | struct | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 16 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 17 | open ATP_Util | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 18 | open ATP_Problem | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 19 | open ATP_Proof | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 20 | open ATP_Proof_Reconstruct | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 21 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 22 | val z3_apply_def_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Apply_Def | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 23 | val z3_hypothesis_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Hypothesis | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 24 | val z3_intro_def_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Intro_Def | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 25 | val z3_lemma_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Lemma | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 26 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 27 | fun inline_z3_defs _ [] = [] | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 28 | | inline_z3_defs defs ((name, role, t, rule, deps) :: lines) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 29 | if rule = z3_intro_def_rule then | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 30 | let val def = t |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> swap in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 31 | inline_z3_defs (insert (op =) def defs) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 32 | (map (replace_dependencies_in_line (name, [])) lines) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 33 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 34 | else if rule = z3_apply_def_rule then | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 35 | inline_z3_defs defs (map (replace_dependencies_in_line (name, [])) lines) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 36 | else | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 37 | (name, role, Term.subst_atomic defs t, rule, deps) :: inline_z3_defs defs lines | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 38 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 39 | fun add_z3_hypotheses [] = I | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 40 | | add_z3_hypotheses hyps = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 41 | HOLogic.dest_Trueprop | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 42 | #> curry s_imp (Library.foldr1 s_conj (map HOLogic.dest_Trueprop hyps)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 43 | #> HOLogic.mk_Trueprop | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 44 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 45 | fun inline_z3_hypotheses _ _ [] = [] | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 46 | | inline_z3_hypotheses hyp_names hyps ((name, role, t, rule, deps) :: lines) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 47 | if rule = z3_hypothesis_rule then | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 48 | inline_z3_hypotheses (name :: hyp_names) (AList.map_default (op =) (t, []) (cons name) hyps) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 49 | lines | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 50 | else | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 51 | let val deps' = subtract (op =) hyp_names deps in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 52 | if rule = z3_lemma_rule then | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 53 | (name, role, t, rule, deps') :: inline_z3_hypotheses hyp_names hyps lines | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 54 | else | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 55 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 56 | val add_hyps = filter_out (null o inter (op =) deps o snd) hyps | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 57 | val t' = add_z3_hypotheses (map fst add_hyps) t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 58 | val deps' = subtract (op =) hyp_names deps | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 59 | val hyps' = fold (AList.update (op =) o apsnd (insert (op =) name)) add_hyps hyps | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 60 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 61 | (name, role, t', rule, deps') :: inline_z3_hypotheses hyp_names hyps' lines | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 63 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 64 | |
| 56855 
e7a55d295b8e
simplify unused universally quantified variables in Z3 proofs
 blanchet parents: 
56129diff
changeset | 65 | fun simplify_bool ((all as Const (@{const_name All}, _)) $ Abs (s, T, t)) =
 | 
| 
e7a55d295b8e
simplify unused universally quantified variables in Z3 proofs
 blanchet parents: 
56129diff
changeset | 66 | let val t' = simplify_bool t in | 
| 
e7a55d295b8e
simplify unused universally quantified variables in Z3 proofs
 blanchet parents: 
56129diff
changeset | 67 | if loose_bvar1 (t', 0) then all $ Abs (s, T, t') else t' | 
| 
e7a55d295b8e
simplify unused universally quantified variables in Z3 proofs
 blanchet parents: 
56129diff
changeset | 68 | end | 
| 
e7a55d295b8e
simplify unused universally quantified variables in Z3 proofs
 blanchet parents: 
56129diff
changeset | 69 |   | simplify_bool (@{const Not} $ t) = s_not (simplify_bool t)
 | 
| 56129 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 70 |   | simplify_bool (@{const conj} $ t $ u) = s_conj (simplify_bool t, simplify_bool u)
 | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 71 |   | simplify_bool (@{const disj} $ t $ u) = s_disj (simplify_bool t, simplify_bool u)
 | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 72 |   | simplify_bool (@{const implies} $ t $ u) = s_imp (simplify_bool t, simplify_bool u)
 | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 73 |   | simplify_bool (@{const HOL.eq (bool)} $ t $ u) = s_iff (simplify_bool t, simplify_bool u)
 | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 74 |   | simplify_bool (t as Const (@{const_name HOL.eq}, _) $ u $ v) =
 | 
| 56126 | 75 |     if u aconv v then @{const True} else t
 | 
| 56129 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 76 | | simplify_bool (t $ u) = simplify_bool t $ simplify_bool u | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 77 | | simplify_bool (Abs (s, T, t)) = Abs (s, T, simplify_bool t) | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 78 | | simplify_bool t = t | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 79 | |
| 56129 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 80 | (* It is not entirely clear why this should be necessary, especially for abstractions variables. *) | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 81 | val unskolemize_names = | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 82 | Term.map_abs_vars (perhaps (try Name.dest_skolem)) | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 83 | #> Term.map_aterms (perhaps (try (fn Free (s, T) => Free (Name.dest_skolem s, T)))) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 | |
| 57289 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 85 | fun strip_alls (Const (@{const_name All}, _) $ Abs (s, T, body)) = strip_alls body |>> cons (s, T)
 | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 86 | | strip_alls t = ([], t) | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 87 | |
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 88 | fun push_skolem_all_under_iff t = | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 89 | (case strip_alls t of | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 90 | (qs as _ :: _, | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 91 |      (t0 as Const (@{const_name HOL.eq}, _)) $ (t1 as Const (@{const_name Ex}, _) $ _) $ t2) =>
 | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 92 | t0 $ HOLogic.list_all (qs, t1) $ HOLogic.list_all (qs, t2) | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 93 | | _ => t) | 
| 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 94 | |
| 57056 | 95 | fun atp_proof_of_z3_proof ctxt rewrite_rules hyp_ts concl_t fact_helper_ts prem_ids conjecture_id | 
| 96 | fact_helper_ids proof = | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 97 | let | 
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 98 | val thy = Proof_Context.theory_of ctxt | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 99 | |
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 100 |     fun steps_of (Z3_New_Proof.Z3_Step {id, rule, prems, concl, ...}) =
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 101 | let | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 102 | val sid = string_of_int id | 
| 56128 | 103 | |
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 104 | val concl' = | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 105 | concl | 
| 56128 | 106 | |> Raw_Simplifier.rewrite_term thy rewrite_rules [] | 
| 107 | |> Object_Logic.atomize_term thy | |
| 56129 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 108 | |> simplify_bool | 
| 
9ee083f9da5b
remove '__' skolem suffixes before showing terms to users
 blanchet parents: 
56128diff
changeset | 109 | |> unskolemize_names | 
| 57289 
5483868da0d8
help reconstruction of Z3 skolemization by weakening formulas a bit
 blanchet parents: 
57218diff
changeset | 110 | |> push_skolem_all_under_iff | 
| 56128 | 111 | |> HOLogic.mk_Trueprop | 
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 112 | |
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 113 | fun standard_step role = | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 114 | ((sid, []), role, concl', Z3_New_Proof.string_of_rule rule, | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 115 | map (fn id => (string_of_int id, [])) prems) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 116 | in | 
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 117 | (case rule of | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 118 | Z3_New_Proof.Asserted => | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 119 | let | 
| 57056 | 120 | val ss = the_list (AList.lookup (op =) fact_helper_ids id) | 
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 121 | val name0 = (sid ^ "a", ss) | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 122 | |
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 123 | val (role0, concl0) = | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 124 | (case ss of | 
| 57056 | 125 | [s] => (Axiom, the (AList.lookup (op =) fact_helper_ts s)) | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 126 | | _ => | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 127 | if id = conjecture_id then | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 128 | (Conjecture, concl_t) | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 129 | else | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 130 | (Hypothesis, | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 131 | (case find_index (curry (op =) id) prem_ids of | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 132 | ~1 => concl | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 133 | | i => nth hyp_ts i))) | 
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 134 | |
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 135 | val normalize_prems = | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 136 | SMT2_Normalize.case_bool_entry :: SMT2_Normalize.special_quant_table @ | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 137 | SMT2_Normalize.abs_min_max_table | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 138 | |> map_filter (fn (c, th) => | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 139 | if exists_Const (curry (op =) c o fst) concl0 then | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 140 | let val s = short_thm_name ctxt th in SOME (s, [s]) end | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 141 | else | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 142 | NONE) | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 143 | in | 
| 57218 | 144 | (if role0 = Axiom then [] | 
| 145 | else [(name0, role0, concl0, Z3_New_Proof.string_of_rule rule, [])]) @ | |
| 146 | [((sid, []), Plain, concl', Z3_New_Proof.string_of_rule Z3_New_Proof.Rewrite, | |
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56981diff
changeset | 147 | name0 :: normalize_prems)] | 
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 148 | end | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 149 | | Z3_New_Proof.Rewrite => [standard_step Lemma] | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 150 | | Z3_New_Proof.Rewrite_Star => [standard_step Lemma] | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 151 | | Z3_New_Proof.Skolemize => [standard_step Lemma] | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 152 | | Z3_New_Proof.Th_Lemma _ => [standard_step Lemma] | 
| 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 153 | | _ => [standard_step Plain]) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 154 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 155 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 156 | proof | 
| 56981 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 blanchet parents: 
56855diff
changeset | 157 | |> maps steps_of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 158 | |> inline_z3_defs [] | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 159 | |> inline_z3_hypotheses [] [] | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 160 | end | 
| 
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 | end; |