src/HOL/Tools/SMT2/z3_new_isar.ML
author blanchet
Fri, 14 Mar 2014 11:15:46 +0100
changeset 56128 c106ac2ff76d
parent 56126 fc937e7ef4c6
child 56129 9ee083f9da5b
permissions -rw-r--r--
undo rewrite rules (e.g. for 'fun_app') in Isar
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     9
  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    10
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
    11
  val atp_proof_of_z3_proof: theory -> thm list -> int -> (int * string) list ->
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
    12
    Z3_New_Proof.z3_step list -> (term, string) atp_step list
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    13
end;
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    14
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
structure Z3_New_Isar: Z3_NEW_ISAR =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    16
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    18
open ATP_Util
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
open ATP_Problem
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    20
open ATP_Proof
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    21
open ATP_Proof_Reconstruct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    22
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    23
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
    24
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
    25
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
    26
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
    27
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
fun inline_z3_defs _ [] = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    29
  | 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
    30
    if rule = z3_intro_def_rule then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
      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
    32
        inline_z3_defs (insert (op =) def defs)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
          (map (replace_dependencies_in_line (name, [])) lines)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    34
      end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
    else if rule = z3_apply_def_rule then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    36
      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
    37
    else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    38
      (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
    39
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
fun add_z3_hypotheses [] = I
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    41
  | add_z3_hypotheses hyps =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    42
    HOLogic.dest_Trueprop
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
    #> 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
    44
    #> HOLogic.mk_Trueprop
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 inline_z3_hypotheses _ _ [] = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    47
  | 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
    48
    if rule = z3_hypothesis_rule then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    49
      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
    50
        lines
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
    else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
      let val deps' = subtract (op =) hyp_names deps in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
        if rule = z3_lemma_rule then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    54
          (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
    55
        else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    56
          let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
            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
    58
            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
    59
            val deps' = subtract (op =) hyp_names deps
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    60
            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
    61
          in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
            (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
    63
          end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    64
      end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
fun simplify_prop (@{const Not} $ t) = s_not (simplify_prop t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    67
  | simplify_prop (@{const conj} $ t $ u) = s_conj (simplify_prop t, simplify_prop u)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    68
  | simplify_prop (@{const disj} $ t $ u) = s_disj (simplify_prop t, simplify_prop u)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    69
  | simplify_prop (@{const implies} $ t $ u) = s_imp (simplify_prop t, simplify_prop u)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    70
  | simplify_prop (@{const HOL.eq (bool)} $ t $ u) = s_iff (simplify_prop t, simplify_prop u)
56126
fc937e7ef4c6 more simplification of trivial steps
blanchet
parents: 56104
diff changeset
    71
  | simplify_prop (t as Const (@{const_name HOL.eq}, _) $ u $ v) =
fc937e7ef4c6 more simplification of trivial steps
blanchet
parents: 56104
diff changeset
    72
    if u aconv v then @{const True} else t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
  | simplify_prop (t $ u) = simplify_prop t $ simplify_prop u
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    74
  | simplify_prop (Abs (s, T, t)) = Abs (s, T, simplify_prop t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    75
  | simplify_prop t = t
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 simplify_line (name, role, t, rule, deps) = (name, role, simplify_prop t, rule, deps)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    78
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
    79
fun atp_proof_of_z3_proof thy rewrite_rules conjecture_id fact_ids proof =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    80
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    81
    fun step_of (Z3_New_Proof.Z3_Step {id, rule, prems, concl, ...}) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    82
      let
56099
bc036c1cf111 thread through step IDs from Z3 to Sledgehammer
blanchet
parents: 56083
diff changeset
    83
        fun step_name_of id = (string_of_int id, the_list (AList.lookup (op =) fact_ids id))
bc036c1cf111 thread through step IDs from Z3 to Sledgehammer
blanchet
parents: 56083
diff changeset
    84
bc036c1cf111 thread through step IDs from Z3 to Sledgehammer
blanchet
parents: 56083
diff changeset
    85
        val name as (_, ss) = step_name_of id
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
    86
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    87
        val role =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
          (case rule of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    89
            Z3_New_Proof.Asserted =>
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
    90
              if not (null ss) then Axiom
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
    91
              else if id = conjecture_id then Negated_Conjecture
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
    92
              else Hypothesis
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    93
          | Z3_New_Proof.Rewrite => Lemma
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    94
          | Z3_New_Proof.Rewrite_Star => Lemma
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    95
          | Z3_New_Proof.Skolemize => Lemma
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    96
          | Z3_New_Proof.Th_Lemma _ => Lemma
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    97
          | _ => Plain)
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
    98
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
    99
        val concl' = concl
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
   100
          |> Raw_Simplifier.rewrite_term thy rewrite_rules []
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
   101
          |> Object_Logic.atomize_term thy
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
   102
          |> HOLogic.mk_Trueprop
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
      in
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56126
diff changeset
   104
        (name, role, concl', Z3_New_Proof.string_of_rule rule, map step_name_of prems)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
      end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
    proof
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   108
    |> map step_of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   109
    |> inline_z3_defs []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   110
    |> inline_z3_hypotheses [] []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   111
    |> map simplify_line
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   112
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   113
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   114
end;