src/HOL/Tools/SMT/z3_proof_reconstruction.ML
author boehmes
Thu, 16 Dec 2010 13:34:28 +0100
changeset 41197 edab1efe0a70
parent 41172 a17c2d669c40
child 41328 6792a5c92a58
permissions -rw-r--r--
fix lambda-lifting: take level of bound variables into account and also apply bound variables from outer scope
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Tools/SMT/z3_proof_reconstruction.ML
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     3
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     4
Proof reconstruction for proofs found by Z3.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     5
*)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     6
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     7
signature Z3_PROOF_RECONSTRUCTION =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     8
sig
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
     9
  val add_z3_rule: thm -> Context.generic -> Context.generic
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    10
  val reconstruct: Proof.context -> SMT_Translate.recon -> string list ->
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
    11
    int list * thm
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    12
  val setup: theory -> theory
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    13
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    14
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    15
structure Z3_Proof_Reconstruction: Z3_PROOF_RECONSTRUCTION =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    16
struct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    17
41172
a17c2d669c40 the functions term_of and prop_of are also needed in earlier stages, not only for Z3 proof reconstruction, so they really belong in SMT_Utils
boehmes
parents: 41131
diff changeset
    18
structure U = SMT_Utils
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    19
structure P = Z3_Proof_Parser
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    20
structure T = Z3_Proof_Tools
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    21
structure L = Z3_Proof_Literals
40662
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents: 40579
diff changeset
    22
structure M = Z3_Proof_Methods
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    23
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40274
diff changeset
    24
fun z3_exn msg = raise SMT_Failure.SMT (SMT_Failure.Other_Failure
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    25
  ("Z3 proof reconstruction: " ^ msg))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    26
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    27
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    28
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
    29
(* net of schematic rules *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    30
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    31
val z3_ruleN = "z3_rule"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    32
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    33
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    34
  val description = "declaration of Z3 proof rules"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    35
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    36
  val eq = Thm.eq_thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    37
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    38
  structure Z3_Rules = Generic_Data
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    39
  (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    40
    type T = thm Net.net
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    41
    val empty = Net.empty
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    42
    val extend = I
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    43
    val merge = Net.merge eq
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
  )
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    45
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    46
  val prep = `Thm.prop_of o Simplifier.rewrite_rule [L.rewrite_true]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    47
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    48
  fun ins thm net = Net.insert_term eq (prep thm) net handle Net.INSERT => net
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
  fun del thm net = Net.delete_term eq (prep thm) net handle Net.DELETE => net
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    50
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    51
  val add = Thm.declaration_attribute (Z3_Rules.map o ins)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    52
  val del = Thm.declaration_attribute (Z3_Rules.map o del)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    53
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    54
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    55
val add_z3_rule = Z3_Rules.map o ins
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    56
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    57
fun by_schematic_rule ctxt ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    58
  the (T.net_instance (Z3_Rules.get (Context.Proof ctxt)) ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    59
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    60
val z3_rules_setup =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    61
  Attrib.setup (Binding.name z3_ruleN) (Attrib.add_del add del) description #>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39020
diff changeset
    62
  Global_Theory.add_thms_dynamic (Binding.name z3_ruleN, Net.content o Z3_Rules.get)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    63
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    64
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    65
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    66
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    67
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
    68
(* proof tools *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    69
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    70
fun named ctxt name prover ct =
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40274
diff changeset
    71
  let val _ = SMT_Config.trace_msg ctxt I ("Z3: trying " ^ name ^ " ...")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    72
  in prover ct end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    73
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    74
fun NAMED ctxt name tac i st =
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40274
diff changeset
    75
  let val _ = SMT_Config.trace_msg ctxt I ("Z3: trying " ^ name ^ " ...")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    76
  in tac i st end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    77
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    78
fun pretty_goal ctxt thms t =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    79
  [Pretty.block [Pretty.str "proposition: ", Syntax.pretty_term ctxt t]]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    80
  |> not (null thms) ? cons (Pretty.big_list "assumptions:"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    81
       (map (Display.pretty_thm ctxt) thms))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    82
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    83
fun try_apply ctxt thms =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    84
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    85
    fun try_apply_err ct = Pretty.string_of (Pretty.chunks [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    86
      Pretty.big_list ("Z3 found a proof," ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    87
        " but proof reconstruction failed at the following subgoal:")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    88
        (pretty_goal ctxt thms (Thm.term_of ct)),
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    89
      Pretty.str ("Adding a rule to the lemma group " ^ quote z3_ruleN ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    90
        " might solve this problem.")])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    91
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    92
    fun apply [] ct = error (try_apply_err ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    93
      | apply (prover :: provers) ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    94
          (case try prover ct of
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40274
diff changeset
    95
            SOME thm => (SMT_Config.trace_msg ctxt I "Z3: succeeded"; thm)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    96
          | NONE => apply provers ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    97
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    98
  in apply o cons (named ctxt "schematic rules" (by_schematic_rule ctxt)) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    99
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   100
local
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   101
  val rewr_if =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   102
    @{lemma "(if P then Q1 else Q2) = ((P --> Q1) & (~P --> Q2))" by simp}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   103
in
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   104
val simp_fast_tac =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   105
  Simplifier.simp_tac (HOL_ss addsimps [rewr_if])
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   106
  THEN_ALL_NEW Classical.fast_tac HOL_cs
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   107
end
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   108
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   109
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   110
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   111
(* theorems and proofs *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   112
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   113
(** theorem incarnations **)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   114
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   115
datatype theorem =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   116
  Thm of thm | (* theorem without special features *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   117
  MetaEq of thm | (* meta equality "t == s" *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   118
  Literals of thm * L.littab
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   119
    (* "P1 & ... & Pn" and table of all literals P1, ..., Pn *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   120
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   121
fun thm_of (Thm thm) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   122
  | thm_of (MetaEq thm) = thm COMP @{thm meta_eq_to_obj_eq}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   123
  | thm_of (Literals (thm, _)) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   124
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   125
fun meta_eq_of (MetaEq thm) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   126
  | meta_eq_of p = mk_meta_eq (thm_of p)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   127
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   128
fun literals_of (Literals (_, lits)) = lits
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   129
  | literals_of p = L.make_littab [thm_of p]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   130
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   131
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   132
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   133
(** core proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   134
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   135
(* assumption *)
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   136
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   137
local
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   138
  val remove_trigger = mk_meta_eq @{thm SMT.trigger_def}
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   139
  val remove_weight = mk_meta_eq @{thm SMT.weight_def}
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   140
  val remove_fun_app = mk_meta_eq @{thm SMT.fun_app_def}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   141
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   142
  fun rewrite_conv ctxt eqs = Simplifier.full_rewrite
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   143
    (Simplifier.context ctxt Simplifier.empty_ss addsimps eqs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   144
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   145
  val prep_rules = [@{thm Let_def}, remove_trigger, remove_weight,
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   146
    remove_fun_app, L.rewrite_true]
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   147
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   148
  fun rewrite ctxt eqs = Conv.fconv_rule (rewrite_conv ctxt eqs)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   149
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   150
  fun burrow_snd_option f (i, thm) = Option.map (pair i) (f thm)
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   151
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   152
  fun lookup_assm assms_net ct =
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   153
    T.net_instance' burrow_snd_option assms_net ct
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   154
    |> Option.map (fn ithm as (_, thm) => (ithm, Thm.cprop_of thm aconvc ct))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   155
in
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   156
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   157
fun add_asserted outer_ctxt rewrite_rules assms asserted ctxt =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   158
  let
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   159
    val eqs = map (rewrite ctxt [L.rewrite_true]) rewrite_rules
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   160
    val eqs' = union Thm.eq_thm eqs prep_rules
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   161
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   162
    val assms_net =
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   163
      assms
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   164
      |> map (apsnd (rewrite ctxt eqs'))
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   165
      |> map (apsnd (Conv.fconv_rule Thm.eta_conversion))
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   166
      |> T.thm_net_of snd 
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   167
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   168
    fun revert_conv ctxt = rewrite_conv ctxt eqs' then_conv Thm.eta_conversion
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   169
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   170
    fun assume thm ctxt =
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   171
      let
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   172
        val ct = Thm.cprem_of thm 1
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   173
        val (thm', ctxt') = yield_singleton Assumption.add_assumes ct ctxt
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   174
      in (Thm.implies_elim thm thm', ctxt') end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   175
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   176
    fun add (idx, ct) ((is, thms), (ctxt, ptab)) =
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   177
      let
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   178
        val thm1 = 
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   179
          Thm.trivial ct
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   180
          |> Conv.fconv_rule (Conv.arg1_conv (revert_conv outer_ctxt))
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   181
        val thm2 = singleton (Variable.export ctxt outer_ctxt) thm1
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   182
      in
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   183
        (case lookup_assm assms_net (Thm.cprem_of thm2 1) of
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   184
          NONE =>
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   185
            let val (thm, ctxt') = assume thm1 ctxt
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   186
            in ((is, thms), (ctxt', Inttab.update (idx, Thm thm) ptab)) end
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   187
        | SOME ((i, th), exact) =>
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   188
            let
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   189
              val (thm, ctxt') =
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   190
                if exact then (Thm.implies_elim thm1 th, ctxt)
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   191
                else assume thm1 ctxt
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   192
              val thms' = if exact then thms else th :: thms
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   193
            in 
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   194
              ((insert (op =) i is, thms'),
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   195
                (ctxt', Inttab.update (idx, Thm thm) ptab))
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   196
            end)
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   197
      end
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   198
  in fold add asserted (([], []), (ctxt, Inttab.empty)) end
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   199
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   200
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   201
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   202
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   203
(* P = Q ==> P ==> Q   or   P --> Q ==> P ==> Q *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   204
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   205
  val meta_iffD1 = @{lemma "P == Q ==> P ==> (Q::bool)" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   206
  val meta_iffD1_c = T.precompose2 Thm.dest_binop meta_iffD1
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   207
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   208
  val iffD1_c = T.precompose2 (Thm.dest_binop o Thm.dest_arg) @{thm iffD1}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   209
  val mp_c = T.precompose2 (Thm.dest_binop o Thm.dest_arg) @{thm mp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   210
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   211
fun mp (MetaEq thm) p = Thm (Thm.implies_elim (T.compose meta_iffD1_c thm) p)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   212
  | mp p_q p = 
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   213
      let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   214
        val pq = thm_of p_q
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   215
        val thm = T.compose iffD1_c pq handle THM _ => T.compose mp_c pq
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   216
      in Thm (Thm.implies_elim thm p) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   217
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   218
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   219
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   220
(* and_elim:     P1 & ... & Pn ==> Pi *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   221
(* not_or_elim:  ~(P1 | ... | Pn) ==> ~Pi *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   222
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   223
  fun is_sublit conj t = L.exists_lit conj (fn u => u aconv t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   224
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   225
  fun derive conj t lits idx ptab =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   226
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   227
      val lit = the (L.get_first_lit (is_sublit conj t) lits)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   228
      val ls = L.explode conj false false [t] lit
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   229
      val lits' = fold L.insert_lit ls (L.delete_lit lit lits)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   230
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   231
      fun upd thm = Literals (thm_of thm, lits')
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   232
    in (the (L.lookup_lit lits' t), Inttab.map_entry idx upd ptab) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   233
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   234
  fun lit_elim conj (p, idx) ct ptab =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   235
    let val lits = literals_of p
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   236
    in
41172
a17c2d669c40 the functions term_of and prop_of are also needed in earlier stages, not only for Z3 proof reconstruction, so they really belong in SMT_Utils
boehmes
parents: 41131
diff changeset
   237
      (case L.lookup_lit lits (U.term_of ct) of
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   238
        SOME lit => (Thm lit, ptab)
41172
a17c2d669c40 the functions term_of and prop_of are also needed in earlier stages, not only for Z3 proof reconstruction, so they really belong in SMT_Utils
boehmes
parents: 41131
diff changeset
   239
      | NONE => apfst Thm (derive conj (U.term_of ct) lits idx ptab))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   240
    end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   241
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   242
val and_elim = lit_elim true
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   243
val not_or_elim = lit_elim false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   244
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   245
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   246
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   247
(* P1, ..., Pn |- False ==> |- ~P1 | ... | ~Pn *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   248
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   249
  fun step lit thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   250
    Thm.implies_elim (Thm.implies_intr (Thm.cprop_of lit) thm) lit
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   251
  val explode_disj = L.explode false false false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   252
  fun intro hyps thm th = fold step (explode_disj hyps th) thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   253
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   254
  fun dest_ccontr ct = [Thm.dest_arg (Thm.dest_arg (Thm.dest_arg1 ct))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   255
  val ccontr = T.precompose dest_ccontr @{thm ccontr}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   256
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   257
fun lemma thm ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   258
  let
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   259
    val cu = L.negate (Thm.dest_arg ct)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   260
    val hyps = map_filter (try HOLogic.dest_Trueprop) (#hyps (Thm.rep_thm thm))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   261
  in Thm (T.compose ccontr (T.under_assumption (intro hyps thm) cu)) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   262
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   263
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   264
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   265
(* \/{P1, ..., Pn, Q1, ..., Qn}, ~P1, ..., ~Pn ==> \/{Q1, ..., Qn} *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   266
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   267
  val explode_disj = L.explode false true false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   268
  val join_disj = L.join false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   269
  fun unit thm thms th =
41172
a17c2d669c40 the functions term_of and prop_of are also needed in earlier stages, not only for Z3 proof reconstruction, so they really belong in SMT_Utils
boehmes
parents: 41131
diff changeset
   270
    let val t = @{const Not} $ U.prop_of thm and ts = map U.prop_of thms
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   271
    in join_disj (L.make_littab (thms @ explode_disj ts th)) t end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   272
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   273
  fun dest_arg2 ct = Thm.dest_arg (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   274
  fun dest ct = pairself dest_arg2 (Thm.dest_binop ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   275
  val contrapos = T.precompose2 dest @{lemma "(~P ==> ~Q) ==> Q ==> P" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   276
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   277
fun unit_resolution thm thms ct =
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   278
  L.negate (Thm.dest_arg ct)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   279
  |> T.under_assumption (unit thm thms)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   280
  |> Thm o T.discharge thm o T.compose contrapos
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   281
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   282
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   283
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   284
(* P ==> P == True   or   P ==> P == False *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   285
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   286
  val iff1 = @{lemma "P ==> P == (~ False)" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   287
  val iff2 = @{lemma "~P ==> P == False" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   288
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   289
fun iff_true thm = MetaEq (thm COMP iff1)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   290
fun iff_false thm = MetaEq (thm COMP iff2)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   291
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   292
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   293
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   294
(* distributivity of | over & *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   295
fun distributivity ctxt = Thm o try_apply ctxt [] [
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   296
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   297
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   298
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   299
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   300
(* Tseitin-like axioms *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   301
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   302
  val disjI1 = @{lemma "(P ==> Q) ==> ~P | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   303
  val disjI2 = @{lemma "(~P ==> Q) ==> P | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   304
  val disjI3 = @{lemma "(~Q ==> P) ==> P | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   305
  val disjI4 = @{lemma "(Q ==> P) ==> P | ~Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   306
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   307
  fun prove' conj1 conj2 ct2 thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   308
    let val lits = L.true_thm :: L.explode conj1 true (conj1 <> conj2) [] thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   309
    in L.join conj2 (L.make_littab lits) (Thm.term_of ct2) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   310
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   311
  fun prove rule (ct1, conj1) (ct2, conj2) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   312
    T.under_assumption (prove' conj1 conj2 ct2) ct1 COMP rule
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   313
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   314
  fun prove_def_axiom ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   315
    let val (ct1, ct2) = Thm.dest_binop (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   316
    in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   317
      (case Thm.term_of ct1 of
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   318
        @{const Not} $ (@{const HOL.conj} $ _ $ _) =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   319
          prove disjI1 (Thm.dest_arg ct1, true) (ct2, true)
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   320
      | @{const HOL.conj} $ _ $ _ =>
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   321
          prove disjI3 (L.negate ct2, false) (ct1, true)
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   322
      | @{const Not} $ (@{const HOL.disj} $ _ $ _) =>
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   323
          prove disjI3 (L.negate ct2, false) (ct1, false)
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   324
      | @{const HOL.disj} $ _ $ _ =>
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   325
          prove disjI2 (L.negate ct1, false) (ct2, true)
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40680
diff changeset
   326
      | Const (@{const_name distinct}, _) $ _ =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   327
          let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   328
            fun dis_conv cv = Conv.arg_conv (Conv.arg1_conv cv)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   329
            fun prv cu =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   330
              let val (cu1, cu2) = Thm.dest_binop (Thm.dest_arg cu)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   331
              in prove disjI4 (Thm.dest_arg cu2, true) (cu1, true) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   332
          in T.with_conv (dis_conv T.unfold_distinct_conv) prv ct end
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40680
diff changeset
   333
      | @{const Not} $ (Const (@{const_name distinct}, _) $ _) =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   334
          let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   335
            fun dis_conv cv = Conv.arg_conv (Conv.arg1_conv (Conv.arg_conv cv))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   336
            fun prv cu =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   337
              let val (cu1, cu2) = Thm.dest_binop (Thm.dest_arg cu)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   338
              in prove disjI1 (Thm.dest_arg cu1, true) (cu2, true) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   339
          in T.with_conv (dis_conv T.unfold_distinct_conv) prv ct end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   340
      | _ => raise CTERM ("prove_def_axiom", [ct]))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   341
    end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   342
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   343
fun def_axiom ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   344
  named ctxt "conj/disj/distinct" prove_def_axiom,
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   345
  T.by_abstraction (true, false) ctxt [] (fn ctxt' =>
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   346
    named ctxt' "simp+fast" (T.by_tac simp_fast_tac))]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   347
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   348
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   349
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   350
(* local definitions *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   351
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   352
  val intro_rules = [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   353
    @{lemma "n == P ==> (~n | P) & (n | ~P)" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   354
    @{lemma "n == (if P then s else t) ==> (~P | n = s) & (P | n = t)"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   355
      by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   356
    @{lemma "n == P ==> n = P" by (rule meta_eq_to_obj_eq)} ]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   357
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   358
  val apply_rules = [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   359
    @{lemma "(~n | P) & (n | ~P) ==> P == n" by (atomize(full)) fast},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   360
    @{lemma "(~P | n = s) & (P | n = t) ==> (if P then s else t) == n"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   361
      by (atomize(full)) fastsimp} ]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   362
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   363
  val inst_rule = T.match_instantiate Thm.dest_arg
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   364
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   365
  fun apply_rule ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   366
    (case get_first (try (inst_rule ct)) intro_rules of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   367
      SOME thm => thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   368
    | NONE => raise CTERM ("intro_def", [ct]))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   369
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   370
fun intro_def ct = T.make_hyp_def (apply_rule ct) #>> Thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   371
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   372
fun apply_def thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   373
  get_first (try (fn rule => MetaEq (thm COMP rule))) apply_rules
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   374
  |> the_default (Thm thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   375
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   376
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   377
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   378
(* negation normal form *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   379
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   380
  val quant_rules1 = ([
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   381
    @{lemma "(!!x. P x == Q) ==> ALL x. P x == Q" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   382
    @{lemma "(!!x. P x == Q) ==> EX x. P x == Q" by simp}], [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   383
    @{lemma "(!!x. P x == Q x) ==> ALL x. P x == ALL x. Q x" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   384
    @{lemma "(!!x. P x == Q x) ==> EX x. P x == EX x. Q x" by simp}])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   385
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   386
  val quant_rules2 = ([
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   387
    @{lemma "(!!x. ~P x == Q) ==> ~(ALL x. P x) == Q" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   388
    @{lemma "(!!x. ~P x == Q) ==> ~(EX x. P x) == Q" by simp}], [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   389
    @{lemma "(!!x. ~P x == Q x) ==> ~(ALL x. P x) == EX x. Q x" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   390
    @{lemma "(!!x. ~P x == Q x) ==> ~(EX x. P x) == ALL x. Q x" by simp}])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   391
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   392
  fun nnf_quant_tac thm (qs as (qs1, qs2)) i st = (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   393
    Tactic.rtac thm ORELSE'
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   394
    (Tactic.match_tac qs1 THEN' nnf_quant_tac thm qs) ORELSE'
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   395
    (Tactic.match_tac qs2 THEN' nnf_quant_tac thm qs)) i st
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   396
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   397
  fun nnf_quant vars qs p ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   398
    T.as_meta_eq ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   399
    |> T.by_tac (nnf_quant_tac (T.varify vars (meta_eq_of p)) qs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   400
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   401
  fun prove_nnf ctxt = try_apply ctxt [] [
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   402
    named ctxt "conj/disj" L.prove_conj_disj_eq,
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   403
    T.by_abstraction (true, false) ctxt [] (fn ctxt' =>
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   404
      named ctxt' "simp+fast" (T.by_tac simp_fast_tac))]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   405
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   406
fun nnf ctxt vars ps ct =
41172
a17c2d669c40 the functions term_of and prop_of are also needed in earlier stages, not only for Z3 proof reconstruction, so they really belong in SMT_Utils
boehmes
parents: 41131
diff changeset
   407
  (case U.term_of ct of
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   408
    _ $ (l as Const _ $ Abs _) $ (r as Const _ $ Abs _) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   409
      if l aconv r
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   410
      then MetaEq (Thm.reflexive (Thm.dest_arg (Thm.dest_arg ct)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   411
      else MetaEq (nnf_quant vars quant_rules1 (hd ps) ct)
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   412
  | _ $ (@{const Not} $ (Const _ $ Abs _)) $ (Const _ $ Abs _) =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   413
      MetaEq (nnf_quant vars quant_rules2 (hd ps) ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   414
  | _ =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   415
      let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   416
        val nnf_rewr_conv = Conv.arg_conv (Conv.arg_conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   417
          (T.unfold_eqs ctxt (map (Thm.symmetric o meta_eq_of) ps)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   418
      in Thm (T.with_conv nnf_rewr_conv (prove_nnf ctxt) ct) end)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   419
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   420
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   421
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   422
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   423
(** equality proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   424
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   425
(* |- t = t *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   426
fun refl ct = MetaEq (Thm.reflexive (Thm.dest_arg (Thm.dest_arg ct)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   427
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   428
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   429
(* s = t ==> t = s *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   430
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   431
  val symm_rule = @{lemma "s = t ==> t == s" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   432
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   433
fun symm (MetaEq thm) = MetaEq (Thm.symmetric thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   434
  | symm p = MetaEq (thm_of p COMP symm_rule)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   435
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   436
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   437
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   438
(* s = t ==> t = u ==> s = u *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   439
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   440
  val trans1 = @{lemma "s == t ==> t =  u ==> s == u" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   441
  val trans2 = @{lemma "s =  t ==> t == u ==> s == u" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   442
  val trans3 = @{lemma "s =  t ==> t =  u ==> s == u" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   443
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   444
fun trans (MetaEq thm1) (MetaEq thm2) = MetaEq (Thm.transitive thm1 thm2)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   445
  | trans (MetaEq thm) q = MetaEq (thm_of q COMP (thm COMP trans1))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   446
  | trans p (MetaEq thm) = MetaEq (thm COMP (thm_of p COMP trans2))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   447
  | trans p q = MetaEq (thm_of q COMP (thm_of p COMP trans3))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   448
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   449
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   450
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   451
(* t1 = s1 ==> ... ==> tn = sn ==> f t1 ... tn = f s1 .. sn
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   452
   (reflexive antecendents are droppped) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   453
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   454
  exception MONO
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   455
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   456
  fun prove_refl (ct, _) = Thm.reflexive ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   457
  fun prove_comb f g cp =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   458
    let val ((ct1, ct2), (cu1, cu2)) = pairself Thm.dest_comb cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   459
    in Thm.combination (f (ct1, cu1)) (g (ct2, cu2)) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   460
  fun prove_arg f = prove_comb prove_refl f
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   461
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   462
  fun prove f cp = prove_comb (prove f) f cp handle CTERM _ => prove_refl cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   463
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   464
  fun prove_nary is_comb f =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   465
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   466
      fun prove (cp as (ct, _)) = f cp handle MONO =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   467
        if is_comb (Thm.term_of ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   468
        then prove_comb (prove_arg prove) prove cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   469
        else prove_refl cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   470
    in prove end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   471
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   472
  fun prove_list f n cp =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   473
    if n = 0 then prove_refl cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   474
    else prove_comb (prove_arg f) (prove_list f (n-1)) cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   475
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   476
  fun with_length f (cp as (cl, _)) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   477
    f (length (HOLogic.dest_list (Thm.term_of cl))) cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   478
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   479
  fun prove_distinct f = prove_arg (with_length (prove_list f))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   480
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   481
  fun prove_eq exn lookup cp =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   482
    (case lookup (Logic.mk_equals (pairself Thm.term_of cp)) of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   483
      SOME eq => eq
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   484
    | NONE => if exn then raise MONO else prove_refl cp)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   485
  
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   486
  val prove_eq_exn = prove_eq true
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   487
  and prove_eq_safe = prove_eq false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   488
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   489
  fun mono f (cp as (cl, _)) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   490
    (case Term.head_of (Thm.term_of cl) of
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   491
      @{const HOL.conj} => prove_nary L.is_conj (prove_eq_exn f)
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   492
    | @{const HOL.disj} => prove_nary L.is_disj (prove_eq_exn f)
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40680
diff changeset
   493
    | Const (@{const_name distinct}, _) => prove_distinct (prove_eq_safe f)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   494
    | _ => prove (prove_eq_safe f)) cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   495
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   496
fun monotonicity eqs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   497
  let
40680
02caa72cb950 be more liberal in reconstructing congruences from Z3 (sometimes the symmetric version of a premise is needed)
boehmes
parents: 40664
diff changeset
   498
    fun and_symmetric (t, thm) = [(t, thm), (t, Thm.symmetric thm)]
02caa72cb950 be more liberal in reconstructing congruences from Z3 (sometimes the symmetric version of a premise is needed)
boehmes
parents: 40664
diff changeset
   499
    val teqs = maps (and_symmetric o `Thm.prop_of o meta_eq_of) eqs
02caa72cb950 be more liberal in reconstructing congruences from Z3 (sometimes the symmetric version of a premise is needed)
boehmes
parents: 40664
diff changeset
   500
    val lookup = AList.lookup (op aconv) teqs
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   501
    val cp = Thm.dest_binop (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   502
  in MetaEq (prove_eq_exn lookup cp handle MONO => mono lookup cp) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   503
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   504
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   505
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   506
(* |- f a b = f b a (where f is equality) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   507
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   508
  val rule = @{lemma "a = b == b = a" by (atomize(full)) (rule eq_commute)}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   509
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   510
fun commutativity ct = MetaEq (T.match_instantiate I (T.as_meta_eq ct) rule)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   511
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   512
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   513
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   514
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   515
(** quantifier proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   516
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   517
(* P ?x = Q ?x ==> (ALL x. P x) = (ALL x. Q x)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   518
   P ?x = Q ?x ==> (EX x. P x) = (EX x. Q x)    *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   519
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   520
  val rules = [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   521
    @{lemma "(!!x. P x == Q x) ==> (ALL x. P x) == (ALL x. Q x)" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   522
    @{lemma "(!!x. P x == Q x) ==> (EX x. P x) == (EX x. Q x)" by simp}]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   523
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   524
fun quant_intro vars p ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   525
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   526
    val thm = meta_eq_of p
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   527
    val rules' = T.varify vars thm :: rules
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   528
    val cu = T.as_meta_eq ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   529
  in MetaEq (T.by_tac (REPEAT_ALL_NEW (Tactic.match_tac rules')) cu) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   530
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   531
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   532
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   533
(* |- ((ALL x. P x) | Q) = (ALL x. P x | Q) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   534
fun pull_quant ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   535
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   536
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   537
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   538
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   539
(* |- (ALL x. P x & Q x) = ((ALL x. P x) & (ALL x. Q x)) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   540
fun push_quant ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   541
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   542
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   543
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   544
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   545
(* |- (ALL x1 ... xn y1 ... yn. P x1 ... xn) = (ALL x1 ... xn. P x1 ... xn) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   546
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   547
  val elim_all = @{lemma "(ALL x. P) == P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   548
  val elim_ex = @{lemma "(EX x. P) == P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   549
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   550
  fun elim_unused_conv ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   551
    Conv.params_conv ~1 (K (Conv.arg_conv (Conv.arg1_conv
36936
c52d1c130898 incorporated further conversions and conversionals, after some minor tuning;
wenzelm
parents: 36899
diff changeset
   552
      (Conv.rewrs_conv [elim_all, elim_ex])))) ctxt
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   553
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   554
  fun elim_unused_tac ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   555
    REPEAT_ALL_NEW (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   556
      Tactic.match_tac [@{thm refl}, @{thm iff_allI}, @{thm iff_exI}]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   557
      ORELSE' CONVERSION (elim_unused_conv ctxt))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   558
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   559
fun elim_unused_vars ctxt = Thm o T.by_tac (elim_unused_tac ctxt)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   560
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   561
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   562
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   563
(* |- (ALL x1 ... xn. ~(x1 = t1 & ... xn = tn) | P x1 ... xn) = P t1 ... tn *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   564
fun dest_eq_res ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   565
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   566
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   567
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   568
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   569
(* |- ~(ALL x1...xn. P x1...xn) | P a1...an *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   570
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   571
  val rule = @{lemma "~ P x | Q ==> ~(ALL x. P x) | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   572
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   573
val quant_inst = Thm o T.by_tac (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   574
  REPEAT_ALL_NEW (Tactic.match_tac [rule])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   575
  THEN' Tactic.rtac @{thm excluded_middle})
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   576
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   577
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   578
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   579
(* c = SOME x. P x |- (EX x. P x) = P c
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   580
   c = SOME x. ~ P x |- ~(ALL x. P x) = ~ P c *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   581
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   582
  val elim_ex = @{lemma "EX x. P == P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   583
  val elim_all = @{lemma "~ (ALL x. P) == ~P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   584
  val sk_ex = @{lemma "c == SOME x. P x ==> EX x. P x == P c"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   585
    by simp (intro eq_reflection some_eq_ex[symmetric])}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   586
  val sk_all = @{lemma "c == SOME x. ~ P x ==> ~(ALL x. P x) == ~ P c"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   587
    by (simp only: not_all) (intro eq_reflection some_eq_ex[symmetric])}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   588
  val sk_ex_rule = ((sk_ex, I), elim_ex)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   589
  and sk_all_rule = ((sk_all, Thm.dest_arg), elim_all)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   590
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   591
  fun dest f sk_rule = 
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   592
    Thm.dest_comb (f (Thm.dest_arg (Thm.dest_arg (Thm.cprop_of sk_rule))))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   593
  fun type_of f sk_rule = Thm.ctyp_of_term (snd (dest f sk_rule))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   594
  fun pair2 (a, b) (c, d) = [(a, c), (b, d)]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   595
  fun inst_sk (sk_rule, f) p c =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   596
    Thm.instantiate ([(type_of f sk_rule, Thm.ctyp_of_term c)], []) sk_rule
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   597
    |> (fn sk' => Thm.instantiate ([], (pair2 (dest f sk') (p, c))) sk')
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   598
    |> Conv.fconv_rule (Thm.beta_conversion true)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   599
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   600
  fun kind (Const (@{const_name Ex}, _) $ _) = (sk_ex_rule, I, I)
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   601
    | kind (@{const Not} $ (Const (@{const_name All}, _) $ _)) =
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40516
diff changeset
   602
        (sk_all_rule, Thm.dest_arg, L.negate)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   603
    | kind t = raise TERM ("skolemize", [t])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   604
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   605
  fun dest_abs_type (Abs (_, T, _)) = T
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   606
    | dest_abs_type t = raise TERM ("dest_abs_type", [t])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   607
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   608
  fun bodies_of thy lhs rhs =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   609
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   610
      val (rule, dest, make) = kind (Thm.term_of lhs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   611
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   612
      fun dest_body idx cbs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   613
        let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   614
          val cb = Thm.dest_arg (dest ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   615
          val T = dest_abs_type (Thm.term_of cb)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   616
          val cv = Thm.cterm_of thy (Var (("x", idx), T))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   617
          val cu = make (Drule.beta_conv cb cv)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   618
          val cbs' = (cv, cb) :: cbs
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   619
        in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   620
          (snd (Thm.first_order_match (cu, rhs)), rev cbs')
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   621
          handle Pattern.MATCH => dest_body (idx+1) cbs' cu
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   622
        end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   623
    in (rule, dest_body 1 [] lhs) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   624
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   625
  fun transitive f thm = Thm.transitive thm (f (Thm.rhs_of thm))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   626
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   627
  fun sk_step (rule, elim) (cv, mct, cb) ((is, thm), ctxt) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   628
    (case mct of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   629
      SOME ct =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   630
        ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   631
        |> T.make_hyp_def (inst_sk rule (Thm.instantiate_cterm ([], is) cb) ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   632
        |>> pair ((cv, ct) :: is) o Thm.transitive thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   633
    | NONE => ((is, transitive (Conv.rewr_conv elim) thm), ctxt))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   634
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   635
fun skolemize ct ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   636
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   637
    val (lhs, rhs) = Thm.dest_binop (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   638
    val (rule, (ctab, cbs)) = bodies_of (ProofContext.theory_of ctxt) lhs rhs
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   639
    fun lookup_var (cv, cb) = (cv, AList.lookup (op aconvc) ctab cv, cb)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   640
  in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   641
    (([], Thm.reflexive lhs), ctxt)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   642
    |> fold (sk_step rule) (map lookup_var cbs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   643
    |>> MetaEq o snd
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   644
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   645
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   646
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   647
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   648
(** theory proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   649
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   650
(* theory lemmas: linear arithmetic, arrays *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   651
fun th_lemma ctxt simpset thms = Thm o try_apply ctxt thms [
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   652
  T.by_abstraction (false, true) ctxt thms (fn ctxt' => T.by_tac (
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   653
    NAMED ctxt' "arith" (Arith_Data.arith_tac ctxt')
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   654
    ORELSE' NAMED ctxt' "simp+arith" (Simplifier.simp_tac simpset THEN_ALL_NEW
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   655
      Arith_Data.arith_tac ctxt')))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   656
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   657
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   658
(* rewriting: prove equalities:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   659
     * ACI of conjunction/disjunction
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   660
     * contradiction, excluded middle
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   661
     * logical rewriting rules (for negation, implication, equivalence,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   662
         distinct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   663
     * normal forms for polynoms (integer/real arithmetic)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   664
     * quantifier elimination over linear arithmetic
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   665
     * ... ? **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   666
structure Z3_Simps = Named_Thms
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   667
(
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   668
  val name = "z3_simp"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   669
  val description = "simplification rules for Z3 proof reconstruction"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   670
)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   671
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   672
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   673
  fun spec_meta_eq_of thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   674
    (case try (fn th => th RS @{thm spec}) thm of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   675
      SOME thm' => spec_meta_eq_of thm'
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   676
    | NONE => mk_meta_eq thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   677
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   678
  fun prep (Thm thm) = spec_meta_eq_of thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   679
    | prep (MetaEq thm) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   680
    | prep (Literals (thm, _)) = spec_meta_eq_of thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   681
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   682
  fun unfold_conv ctxt ths =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   683
    Conv.arg_conv (Conv.binop_conv (T.unfold_eqs ctxt (map prep ths)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   684
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   685
  fun with_conv _ [] prv = prv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   686
    | with_conv ctxt ths prv = T.with_conv (unfold_conv ctxt ths) prv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   687
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   688
  val unfold_conv =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   689
    Conv.arg_conv (Conv.binop_conv (Conv.try_conv T.unfold_distinct_conv))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   690
  val prove_conj_disj_eq = T.with_conv unfold_conv L.prove_conj_disj_eq
40663
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   691
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   692
  fun assume_prems ctxt thm =
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   693
    Assumption.add_assumes (Drule.cprems_of thm) ctxt
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   694
    |>> (fn thms => fold Thm.elim_implies thms thm)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   695
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   696
40663
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   697
fun rewrite simpset ths ct ctxt =
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   698
  apfst Thm (assume_prems ctxt (with_conv ctxt ths (try_apply ctxt [] [
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   699
    named ctxt "conj/disj/distinct" prove_conj_disj_eq,
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   700
    T.by_abstraction (true, false) ctxt [] (fn ctxt' => T.by_tac (
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   701
      NAMED ctxt' "simp (logic)" (Simplifier.simp_tac simpset)
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   702
      THEN_ALL_NEW NAMED ctxt' "fast (logic)" (Classical.fast_tac HOL_cs))),
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   703
    T.by_abstraction (false, true) ctxt [] (fn ctxt' => T.by_tac (
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   704
      NAMED ctxt' "simp (theory)" (Simplifier.simp_tac simpset)
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   705
      THEN_ALL_NEW (
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   706
        NAMED ctxt' "fast (theory)" (Classical.fast_tac HOL_cs)
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   707
        ORELSE' NAMED ctxt' "arith (theory)" (Arith_Data.arith_tac ctxt')))),
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   708
    T.by_abstraction (true, true) ctxt [] (fn ctxt' => T.by_tac (
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   709
      NAMED ctxt' "simp (full)" (Simplifier.simp_tac simpset)
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   710
      THEN_ALL_NEW (
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   711
        NAMED ctxt' "fast (full)" (Classical.fast_tac HOL_cs)
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   712
        ORELSE' NAMED ctxt' "arith (full)" (Arith_Data.arith_tac ctxt')))),
e080c9e68752 share and use more utility functions;
boehmes
parents: 40662
diff changeset
   713
    named ctxt "injectivity" (M.prove_injectivity ctxt)]) ct))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   714
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   715
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   716
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   717
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   718
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   719
(* proof reconstruction *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   720
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   721
(** tracing and checking **)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   722
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   723
fun trace_before ctxt idx = SMT_Config.trace_msg ctxt (fn r =>
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   724
  "Z3: #" ^ string_of_int idx ^ ": " ^ P.string_of_rule r)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   725
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   726
fun check_after idx r ps ct (p, (ctxt, _)) =
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   727
  if not (Config.get ctxt SMT_Config.trace) then ()
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   728
  else
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   729
    let val thm = thm_of p |> tap (Thm.join_proofs o single)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   730
    in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   731
      if (Thm.cprop_of thm) aconvc ct then ()
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   732
      else z3_exn (Pretty.string_of (Pretty.big_list ("proof step failed: " ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   733
        quote (P.string_of_rule r) ^ " (#" ^ string_of_int idx ^ ")")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   734
          (pretty_goal ctxt (map (thm_of o fst) ps) (Thm.prop_of thm) @
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   735
           [Pretty.block [Pretty.str "expected: ",
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   736
            Syntax.pretty_term ctxt (Thm.term_of ct)]])))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   737
    end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   738
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   739
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   740
(** overall reconstruction procedure **)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   741
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   742
local
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   743
  fun not_supported r = raise Fail ("Z3: proof rule not implemented: " ^
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   744
    quote (P.string_of_rule r))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   745
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   746
  fun prove_step simpset vars r ps ct (cxp as (cx, ptab)) =
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   747
    (case (r, ps) of
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   748
      (* core rules *)
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   749
      (P.True_Axiom, _) => (Thm L.true_thm, cxp)
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   750
    | (P.Asserted, _) => raise Fail "bad assertion"
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   751
    | (P.Goal, _) => raise Fail "bad assertion"
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   752
    | (P.Modus_Ponens, [(p, _), (q, _)]) => (mp q (thm_of p), cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   753
    | (P.Modus_Ponens_Oeq, [(p, _), (q, _)]) => (mp q (thm_of p), cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   754
    | (P.And_Elim, [(p, i)]) => and_elim (p, i) ct ptab ||> pair cx
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   755
    | (P.Not_Or_Elim, [(p, i)]) => not_or_elim (p, i) ct ptab ||> pair cx
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   756
    | (P.Hypothesis, _) => (Thm (Thm.assume ct), cxp)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   757
    | (P.Lemma, [(p, _)]) => (lemma (thm_of p) ct, cxp)
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   758
    | (P.Unit_Resolution, (p, _) :: ps) =>
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   759
        (unit_resolution (thm_of p) (map (thm_of o fst) ps) ct, cxp)
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   760
    | (P.Iff_True, [(p, _)]) => (iff_true (thm_of p), cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   761
    | (P.Iff_False, [(p, _)]) => (iff_false (thm_of p), cxp)
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   762
    | (P.Distributivity, _) => (distributivity cx ct, cxp)
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   763
    | (P.Def_Axiom, _) => (def_axiom cx ct, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   764
    | (P.Intro_Def, _) => intro_def ct cx ||> rpair ptab
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   765
    | (P.Apply_Def, [(p, _)]) => (apply_def (thm_of p), cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   766
    | (P.Iff_Oeq, [(p, _)]) => (p, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   767
    | (P.Nnf_Pos, _) => (nnf cx vars (map fst ps) ct, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   768
    | (P.Nnf_Neg, _) => (nnf cx vars (map fst ps) ct, cxp)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   769
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   770
      (* equality rules *)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   771
    | (P.Reflexivity, _) => (refl ct, cxp)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   772
    | (P.Symmetry, [(p, _)]) => (symm p, cxp)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   773
    | (P.Transitivity, [(p, _), (q, _)]) => (trans p q, cxp)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   774
    | (P.Monotonicity, _) => (monotonicity (map fst ps) ct, cxp)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   775
    | (P.Commutativity, _) => (commutativity ct, cxp)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   776
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   777
      (* quantifier rules *)
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   778
    | (P.Quant_Intro, [(p, _)]) => (quant_intro vars p ct, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   779
    | (P.Pull_Quant, _) => (pull_quant cx ct, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   780
    | (P.Push_Quant, _) => (push_quant cx ct, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   781
    | (P.Elim_Unused_Vars, _) => (elim_unused_vars cx ct, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   782
    | (P.Dest_Eq_Res, _) => (dest_eq_res cx ct, cxp)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   783
    | (P.Quant_Inst, _) => (quant_inst ct, cxp)
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   784
    | (P.Skolemize, _) => skolemize ct cx ||> rpair ptab
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   785
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   786
      (* theory rules *)
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   787
    | (P.Th_Lemma _, _) =>  (* FIXME: use arguments *)
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   788
        (th_lemma cx simpset (map (thm_of o fst) ps) ct, cxp)
40662
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents: 40579
diff changeset
   789
    | (P.Rewrite, _) => rewrite simpset [] ct cx ||> rpair ptab
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   790
    | (P.Rewrite_Star, ps) => rewrite simpset (map fst ps) ct cx ||> rpair ptab
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   791
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   792
    | (P.Nnf_Star, _) => not_supported r
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   793
    | (P.Cnf_Star, _) => not_supported r
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   794
    | (P.Transitivity_Star, _) => not_supported r
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   795
    | (P.Pull_Quant_Star, _) => not_supported r
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   796
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   797
    | _ => raise Fail ("Z3: proof rule " ^ quote (P.string_of_rule r) ^
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   798
       " has an unexpected number of arguments."))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   799
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   800
  fun lookup_proof ptab idx =
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   801
    (case Inttab.lookup ptab idx of
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   802
      SOME p => (p, idx)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   803
    | NONE => z3_exn ("unknown proof id: " ^ quote (string_of_int idx)))
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   804
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   805
  fun prove simpset vars (idx, step) (_, cxp as (ctxt, ptab)) =
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   806
    let
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   807
      val P.Proof_Step {rule=r, prems, prop, ...} = step
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   808
      val ps = map (lookup_proof ptab) prems
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   809
      val _ = trace_before ctxt idx r
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   810
      val (thm, (ctxt', ptab')) =
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   811
        cxp
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   812
        |> prove_step simpset vars r ps prop
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   813
        |> tap (check_after idx r ps prop)
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   814
    in (thm, (ctxt', Inttab.update (idx, thm) ptab')) end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   815
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   816
  val disch_rules = [@{thm allI}, @{thm refl}, @{thm reflexive}]
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   817
  fun all_disch_rules rules = map (pair false) (disch_rules @ rules)
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   818
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   819
  fun disch_assm rules thm =
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   820
    if Thm.nprems_of thm = 0 then Drule.flexflex_unique thm
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   821
    else
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   822
      (case Seq.pull (Thm.biresolution false rules 1 thm) of
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   823
        SOME (thm', _) => disch_assm rules thm'
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   824
      | NONE => raise THM ("failed to discharge premise", 1, [thm]))
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   825
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   826
  fun discharge rules outer_ctxt (p, (inner_ctxt, _)) =
41130
130771a48c70 adapted the Z3 proof parser to recent changes in Z3's proof format;
boehmes
parents: 41127
diff changeset
   827
    thm_of p
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   828
    |> singleton (ProofContext.export inner_ctxt outer_ctxt)
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   829
    |> disch_assm rules
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   830
in
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   831
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   832
fun reconstruct outer_ctxt recon output =
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   833
  let
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   834
    val {context=ctxt, typs, terms, rewrite_rules, assms} = recon
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   835
    val (asserted, steps, vars, ctxt1) = P.parse ctxt typs terms output
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   836
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   837
    val simpset = T.make_simpset ctxt1 (Z3_Simps.get ctxt1)
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   838
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   839
    val ((is, rules), cxp as (ctxt2, _)) =
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   840
      add_asserted outer_ctxt rewrite_rules assms asserted ctxt1
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   841
  in
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   842
    if Config.get ctxt2 SMT_Config.filter_only_facts then (is, @{thm TrueI})
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   843
    else
41131
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   844
      (Thm @{thm TrueI}, cxp)
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   845
      |> fold (prove simpset vars) steps 
fc9d503c3d67 fixed proof reconstruction for lambda-lifted problems (which broke when the lambda-lifting code was changed to operate on terms instead of on theorems)
boehmes
parents: 41130
diff changeset
   846
      |> discharge (all_disch_rules rules) outer_ctxt
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 40681
diff changeset
   847
      |> pair []
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   848
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   849
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   850
end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   851
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   852
val setup = z3_rules_setup #> Z3_Simps.setup
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   853
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   854
end