src/HOL/Tools/SMT/z3_proof_reconstruction.ML
author bulwahn
Mon, 22 Nov 2010 11:35:09 +0100
changeset 40658 5ccfc3ee7fe6
parent 40579 98ebd2300823
child 40662 798aad2229c0
permissions -rw-r--r--
adapting example in Predicate_Compile_Examples
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 ->
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39557
diff changeset
    11
    (int list * thm) * Proof.context
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
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    18
structure P = Z3_Proof_Parser
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    19
structure T = Z3_Proof_Tools
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    20
structure L = Z3_Proof_Literals
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    21
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
    22
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
    23
  ("Z3 proof reconstruction: " ^ msg))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    24
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    25
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    26
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    27
(** net of schematic rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    28
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    29
val z3_ruleN = "z3_rule"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    30
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    31
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    32
  val description = "declaration of Z3 proof rules"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    33
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    34
  val eq = Thm.eq_thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    35
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    36
  structure Z3_Rules = Generic_Data
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    37
  (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    38
    type T = thm Net.net
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    39
    val empty = Net.empty
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    40
    val extend = I
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    41
    val merge = Net.merge eq
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    42
  )
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    43
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
  val prep = `Thm.prop_of o Simplifier.rewrite_rule [L.rewrite_true]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    45
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    46
  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
    47
  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
    48
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
  val add = Thm.declaration_attribute (Z3_Rules.map o ins)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    50
  val del = Thm.declaration_attribute (Z3_Rules.map o del)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    51
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    52
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    53
val add_z3_rule = Z3_Rules.map o ins
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    54
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    55
fun by_schematic_rule ctxt ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    56
  the (T.net_instance (Z3_Rules.get (Context.Proof ctxt)) ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    57
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    58
val z3_rules_setup =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    59
  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
    60
  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
    61
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    62
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    63
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    64
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    65
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    66
(** proof tools **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    67
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    68
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
    69
  let val _ = SMT_Config.trace_msg ctxt I ("Z3: trying " ^ name ^ " ...")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    70
  in prover ct end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    71
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    72
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
    73
  let val _ = SMT_Config.trace_msg ctxt I ("Z3: trying " ^ name ^ " ...")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    74
  in tac i st end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    75
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    76
fun pretty_goal ctxt thms t =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    77
  [Pretty.block [Pretty.str "proposition: ", Syntax.pretty_term ctxt t]]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    78
  |> not (null thms) ? cons (Pretty.big_list "assumptions:"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    79
       (map (Display.pretty_thm ctxt) thms))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    80
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    81
fun try_apply ctxt thms =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    82
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    83
    fun try_apply_err ct = Pretty.string_of (Pretty.chunks [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    84
      Pretty.big_list ("Z3 found a proof," ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    85
        " but proof reconstruction failed at the following subgoal:")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    86
        (pretty_goal ctxt thms (Thm.term_of ct)),
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    87
      Pretty.str ("Adding a rule to the lemma group " ^ quote z3_ruleN ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    88
        " might solve this problem.")])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    89
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    90
    fun apply [] ct = error (try_apply_err ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    91
      | apply (prover :: provers) ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    92
          (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
    93
            SOME thm => (SMT_Config.trace_msg ctxt I "Z3: succeeded"; thm)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    94
          | NONE => apply provers ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    95
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    96
  in apply o cons (named ctxt "schematic rules" (by_schematic_rule ctxt)) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    97
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    98
local
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    99
  val rewr_if =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   100
    @{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
   101
in
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   102
val simp_fast_tac =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   103
  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
   104
  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
   105
end
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   106
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   107
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   108
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   109
(** theorems and proofs **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   110
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   111
(* theorem incarnations *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   112
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   113
datatype theorem =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   114
  Thm of thm | (* theorem without special features *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   115
  MetaEq of thm | (* meta equality "t == s" *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   116
  Literals of thm * L.littab
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   117
    (* "P1 & ... & Pn" and table of all literals P1, ..., Pn *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   118
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   119
fun thm_of (Thm thm) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   120
  | thm_of (MetaEq thm) = thm COMP @{thm meta_eq_to_obj_eq}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   121
  | thm_of (Literals (thm, _)) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   122
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   123
fun meta_eq_of (MetaEq thm) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   124
  | meta_eq_of p = mk_meta_eq (thm_of p)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   125
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   126
fun literals_of (Literals (_, lits)) = lits
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   127
  | literals_of p = L.make_littab [thm_of p]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   128
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   129
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   130
(* proof representation *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   131
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   132
datatype proof = Unproved of P.proof_step | Proved of theorem
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   133
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   134
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   135
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   136
(** core proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   137
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   138
(* assumption *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   139
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   140
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   141
  val remove_trigger = @{lemma "trigger t p == p"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   142
    by (rule eq_reflection, rule trigger_def)}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   143
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   144
  val prep_rules = [@{thm Let_def}, remove_trigger, L.rewrite_true]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   145
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   146
  fun rewrite_conv ctxt eqs = Simplifier.full_rewrite
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   147
    (Simplifier.context ctxt Simplifier.empty_ss addsimps eqs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   148
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
   149
  fun rewrites f ctxt eqs = map (f (Conv.fconv_rule (rewrite_conv ctxt eqs)))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   150
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
   151
  fun burrow_snd_option f (i, thm) = Option.map (pair i) (f thm)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   152
  fun lookup_assm ctxt assms ct =
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
   153
    (case T.net_instance' burrow_snd_option assms ct 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
   154
      SOME ithm => ithm
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   155
    | _ => z3_exn ("not asserted: " ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   156
        quote (Syntax.string_of_term ctxt (Thm.term_of ct))))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   157
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   158
fun prepare_assms ctxt unfolds assms =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   159
  let
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
   160
    val unfolds' = rewrites I ctxt [L.rewrite_true] unfolds
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
   161
    val assms' =
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
   162
      rewrites apsnd ctxt (union Thm.eq_thm unfolds' prep_rules) assms
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
   163
  in (unfolds', T.thm_net_of snd assms') end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   164
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   165
fun asserted ctxt (unfolds, assms) ct =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   166
  let val revert_conv = rewrite_conv ctxt unfolds
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
   167
  in Thm (T.with_conv revert_conv (snd o lookup_assm ctxt assms) ct) end
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
   168
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
   169
fun find_assm ctxt (unfolds, assms) ct =
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
   170
  fst (lookup_assm ctxt assms (Thm.rhs_of (rewrite_conv ctxt unfolds ct)))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   171
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   172
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   173
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   174
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   175
(* P = Q ==> P ==> Q   or   P --> Q ==> P ==> Q *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   176
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   177
  val meta_iffD1 = @{lemma "P == Q ==> P ==> (Q::bool)" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   178
  val meta_iffD1_c = T.precompose2 Thm.dest_binop meta_iffD1
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   179
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   180
  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
   181
  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
   182
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   183
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
   184
  | mp p_q p = 
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   185
      let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   186
        val pq = thm_of p_q
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   187
        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
   188
      in Thm (Thm.implies_elim thm p) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   189
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   190
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   191
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   192
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   193
(* and_elim:     P1 & ... & Pn ==> Pi *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   194
(* not_or_elim:  ~(P1 | ... | Pn) ==> ~Pi *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   195
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   196
  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
   197
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   198
  fun derive conj t lits idx ptab =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   199
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   200
      val lit = the (L.get_first_lit (is_sublit conj t) lits)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   201
      val ls = L.explode conj false false [t] lit
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   202
      val lits' = fold L.insert_lit ls (L.delete_lit lit lits)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   203
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   204
      fun upd (Proved thm) = Proved (Literals (thm_of thm, lits'))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   205
        | upd p = p
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   206
    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
   207
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   208
  fun lit_elim conj (p, idx) ct ptab =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   209
    let val lits = literals_of p
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
      (case L.lookup_lit lits (T.term_of ct) of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   212
        SOME lit => (Thm lit, ptab)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   213
      | NONE => apfst Thm (derive conj (T.term_of ct) lits idx ptab))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   214
    end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   215
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   216
val and_elim = lit_elim true
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   217
val not_or_elim = lit_elim false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   218
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   219
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   220
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   221
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   222
(* P1, ..., Pn |- False ==> |- ~P1 | ... | ~Pn *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   223
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   224
  fun step lit thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   225
    Thm.implies_elim (Thm.implies_intr (Thm.cprop_of lit) thm) lit
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   226
  val explode_disj = L.explode false false false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   227
  fun intro hyps thm th = fold step (explode_disj hyps th) thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   228
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   229
  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
   230
  val ccontr = T.precompose dest_ccontr @{thm ccontr}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   231
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   232
fun lemma thm ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   233
  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
   234
    val cu = L.negate (Thm.dest_arg ct)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   235
    val hyps = map_filter (try HOLogic.dest_Trueprop) (#hyps (Thm.rep_thm thm))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   236
  in Thm (T.compose ccontr (T.under_assumption (intro hyps thm) cu)) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   237
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   238
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   239
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   240
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   241
(* \/{P1, ..., Pn, Q1, ..., Qn}, ~P1, ..., ~Pn ==> \/{Q1, ..., Qn} *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   242
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   243
  val explode_disj = L.explode false true false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   244
  val join_disj = L.join false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   245
  fun unit thm thms th =
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
   246
    let val t = @{const Not} $ T.prop_of thm and ts = map T.prop_of thms
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   247
    in join_disj (L.make_littab (thms @ explode_disj ts th)) t end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   248
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   249
  fun dest_arg2 ct = Thm.dest_arg (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   250
  fun dest ct = pairself dest_arg2 (Thm.dest_binop ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   251
  val contrapos = T.precompose2 dest @{lemma "(~P ==> ~Q) ==> Q ==> P" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   252
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   253
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
   254
  L.negate (Thm.dest_arg ct)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   255
  |> T.under_assumption (unit thm thms)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   256
  |> Thm o T.discharge thm o T.compose contrapos
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   257
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   258
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   259
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   260
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   261
(* P ==> P == True   or   P ==> P == False *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   262
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   263
  val iff1 = @{lemma "P ==> P == (~ False)" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   264
  val iff2 = @{lemma "~P ==> P == False" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   265
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   266
fun iff_true thm = MetaEq (thm COMP iff1)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   267
fun iff_false thm = MetaEq (thm COMP iff2)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   268
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   269
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   270
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   271
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   272
(* distributivity of | over & *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   273
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
   274
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   275
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   276
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   277
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   278
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   279
(* Tseitin-like axioms *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   280
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   281
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   282
  val disjI1 = @{lemma "(P ==> Q) ==> ~P | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   283
  val disjI2 = @{lemma "(~P ==> Q) ==> P | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   284
  val disjI3 = @{lemma "(~Q ==> P) ==> P | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   285
  val disjI4 = @{lemma "(Q ==> P) ==> P | ~Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   286
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   287
  fun prove' conj1 conj2 ct2 thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   288
    let val lits = L.true_thm :: L.explode conj1 true (conj1 <> conj2) [] thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   289
    in L.join conj2 (L.make_littab lits) (Thm.term_of ct2) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   290
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   291
  fun prove rule (ct1, conj1) (ct2, conj2) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   292
    T.under_assumption (prove' conj1 conj2 ct2) ct1 COMP rule
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   293
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   294
  fun prove_def_axiom ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   295
    let val (ct1, ct2) = Thm.dest_binop (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   296
    in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   297
      (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
   298
        @{const Not} $ (@{const HOL.conj} $ _ $ _) =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   299
          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
   300
      | @{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
   301
          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
   302
      | @{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
   303
          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
   304
      | @{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
   305
          prove disjI2 (L.negate ct1, false) (ct2, true)
40274
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40164
diff changeset
   306
      | Const (@{const_name SMT.distinct}, _) $ _ =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   307
          let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   308
            fun dis_conv cv = Conv.arg_conv (Conv.arg1_conv cv)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   309
            fun prv cu =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   310
              let val (cu1, cu2) = Thm.dest_binop (Thm.dest_arg cu)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   311
              in prove disjI4 (Thm.dest_arg cu2, true) (cu1, true) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   312
          in T.with_conv (dis_conv T.unfold_distinct_conv) prv ct end
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
   313
      | @{const Not} $ (Const (@{const_name SMT.distinct}, _) $ _) =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   314
          let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   315
            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
   316
            fun prv cu =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   317
              let val (cu1, cu2) = Thm.dest_binop (Thm.dest_arg cu)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   318
              in prove disjI1 (Thm.dest_arg cu1, true) (cu2, true) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   319
          in T.with_conv (dis_conv T.unfold_distinct_conv) prv ct end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   320
      | _ => raise CTERM ("prove_def_axiom", [ct]))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   321
    end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   322
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   323
fun def_axiom ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   324
  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
   325
  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
   326
    named ctxt' "simp+fast" (T.by_tac simp_fast_tac))]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   327
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   328
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   329
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   330
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   331
(* local definitions *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   332
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   333
  val intro_rules = [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   334
    @{lemma "n == P ==> (~n | P) & (n | ~P)" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   335
    @{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
   336
      by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   337
    @{lemma "n == P ==> n = P" by (rule meta_eq_to_obj_eq)} ]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   338
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   339
  val apply_rules = [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   340
    @{lemma "(~n | P) & (n | ~P) ==> P == n" by (atomize(full)) fast},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   341
    @{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
   342
      by (atomize(full)) fastsimp} ]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   343
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   344
  val inst_rule = T.match_instantiate Thm.dest_arg
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   345
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   346
  fun apply_rule ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   347
    (case get_first (try (inst_rule ct)) intro_rules of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   348
      SOME thm => thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   349
    | NONE => raise CTERM ("intro_def", [ct]))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   350
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   351
fun intro_def ct = T.make_hyp_def (apply_rule ct) #>> Thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   352
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   353
fun apply_def thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   354
  get_first (try (fn rule => MetaEq (thm COMP rule))) apply_rules
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   355
  |> the_default (Thm thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   356
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   357
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   358
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   359
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   360
(* negation normal form *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   361
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   362
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   363
  val quant_rules1 = ([
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   364
    @{lemma "(!!x. P x == Q) ==> ALL x. P x == Q" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   365
    @{lemma "(!!x. P x == Q) ==> EX x. P x == Q" by simp}], [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   366
    @{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
   367
    @{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
   368
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   369
  val quant_rules2 = ([
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   370
    @{lemma "(!!x. ~P x == Q) ==> ~(ALL x. P x) == Q" by simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   371
    @{lemma "(!!x. ~P x == Q) ==> ~(EX x. P x) == Q" by simp}], [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   372
    @{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
   373
    @{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
   374
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   375
  fun nnf_quant_tac thm (qs as (qs1, qs2)) i st = (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   376
    Tactic.rtac thm ORELSE'
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   377
    (Tactic.match_tac qs1 THEN' nnf_quant_tac thm qs) ORELSE'
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   378
    (Tactic.match_tac qs2 THEN' nnf_quant_tac thm qs)) i st
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   379
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   380
  fun nnf_quant vars qs p ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   381
    T.as_meta_eq ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   382
    |> 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
   383
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   384
  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
   385
    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
   386
    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
   387
      named ctxt' "simp+fast" (T.by_tac simp_fast_tac))]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   388
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   389
fun nnf ctxt vars ps ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   390
  (case T.term_of ct of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   391
    _ $ (l as Const _ $ Abs _) $ (r as Const _ $ Abs _) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   392
      if l aconv r
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   393
      then MetaEq (Thm.reflexive (Thm.dest_arg (Thm.dest_arg ct)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   394
      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
   395
  | _ $ (@{const Not} $ (Const _ $ Abs _)) $ (Const _ $ Abs _) =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   396
      MetaEq (nnf_quant vars quant_rules2 (hd ps) ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   397
  | _ =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   398
      let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   399
        val nnf_rewr_conv = Conv.arg_conv (Conv.arg_conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   400
          (T.unfold_eqs ctxt (map (Thm.symmetric o meta_eq_of) ps)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   401
      in Thm (T.with_conv nnf_rewr_conv (prove_nnf ctxt) ct) end)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   402
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   403
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   404
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   405
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   406
(** equality proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   407
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   408
(* |- t = t *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   409
fun refl ct = MetaEq (Thm.reflexive (Thm.dest_arg (Thm.dest_arg ct)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   410
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   411
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   412
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   413
(* s = t ==> t = s *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   414
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   415
  val symm_rule = @{lemma "s = t ==> t == s" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   416
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   417
fun symm (MetaEq thm) = MetaEq (Thm.symmetric thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   418
  | symm p = MetaEq (thm_of p COMP symm_rule)
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
(* s = t ==> t = u ==> s = u *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   424
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   425
  val trans1 = @{lemma "s == t ==> t =  u ==> s == u" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   426
  val trans2 = @{lemma "s =  t ==> t == u ==> s == u" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   427
  val trans3 = @{lemma "s =  t ==> t =  u ==> s == u" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   428
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   429
fun trans (MetaEq thm1) (MetaEq thm2) = MetaEq (Thm.transitive thm1 thm2)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   430
  | trans (MetaEq thm) q = MetaEq (thm_of q COMP (thm COMP trans1))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   431
  | trans p (MetaEq thm) = MetaEq (thm COMP (thm_of p COMP trans2))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   432
  | trans p q = MetaEq (thm_of q COMP (thm_of p COMP trans3))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   433
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   434
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   435
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   436
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   437
(* t1 = s1 ==> ... ==> tn = sn ==> f t1 ... tn = f s1 .. sn
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   438
   (reflexive antecendents are droppped) *)
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
  exception MONO
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   441
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   442
  fun prove_refl (ct, _) = Thm.reflexive ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   443
  fun prove_comb f g cp =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   444
    let val ((ct1, ct2), (cu1, cu2)) = pairself Thm.dest_comb cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   445
    in Thm.combination (f (ct1, cu1)) (g (ct2, cu2)) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   446
  fun prove_arg f = prove_comb prove_refl f
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   447
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   448
  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
   449
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   450
  fun prove_nary is_comb f =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   451
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   452
      fun prove (cp as (ct, _)) = f cp handle MONO =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   453
        if is_comb (Thm.term_of ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   454
        then prove_comb (prove_arg prove) prove cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   455
        else prove_refl cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   456
    in prove end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   457
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   458
  fun prove_list f n cp =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   459
    if n = 0 then prove_refl cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   460
    else prove_comb (prove_arg f) (prove_list f (n-1)) cp
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 with_length f (cp as (cl, _)) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   463
    f (length (HOLogic.dest_list (Thm.term_of cl))) cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   464
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   465
  fun prove_distinct f = prove_arg (with_length (prove_list f))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   466
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   467
  fun prove_eq exn lookup cp =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   468
    (case lookup (Logic.mk_equals (pairself Thm.term_of cp)) of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   469
      SOME eq => eq
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   470
    | NONE => if exn then raise MONO else prove_refl cp)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   471
  
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   472
  val prove_eq_exn = prove_eq true
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   473
  and prove_eq_safe = prove_eq false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   474
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   475
  fun mono f (cp as (cl, _)) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   476
    (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
   477
      @{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
   478
    | @{const HOL.disj} => prove_nary L.is_disj (prove_eq_exn f)
40274
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40164
diff changeset
   479
    | Const (@{const_name SMT.distinct}, _) => prove_distinct (prove_eq_safe f)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   480
    | _ => prove (prove_eq_safe f)) cp
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   481
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   482
fun monotonicity eqs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   483
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   484
    val lookup = AList.lookup (op aconv) (map (`Thm.prop_of o meta_eq_of) eqs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   485
    val cp = Thm.dest_binop (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   486
  in MetaEq (prove_eq_exn lookup cp handle MONO => mono lookup cp) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   487
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   488
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   489
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   490
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   491
(* |- f a b = f b a (where f is equality) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   492
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   493
  val rule = @{lemma "a = b == b = a" by (atomize(full)) (rule eq_commute)}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   494
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   495
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
   496
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   497
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   498
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   499
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   500
(** quantifier proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   501
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   502
(* P ?x = Q ?x ==> (ALL x. P x) = (ALL x. Q x)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   503
   P ?x = Q ?x ==> (EX x. P x) = (EX x. Q x)    *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   504
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   505
  val rules = [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   506
    @{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
   507
    @{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
   508
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   509
fun quant_intro vars p ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   510
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   511
    val thm = meta_eq_of p
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   512
    val rules' = T.varify vars thm :: rules
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   513
    val cu = T.as_meta_eq ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   514
  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
   515
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   516
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   517
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   518
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   519
(* |- ((ALL x. P x) | Q) = (ALL x. P x | Q) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   520
fun pull_quant ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   521
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   522
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   523
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   524
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   525
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   526
(* |- (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
   527
fun push_quant ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   528
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   529
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   530
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 x1 ... xn y1 ... yn. P x1 ... xn) = (ALL x1 ... xn. P x1 ... xn) *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   534
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   535
  val elim_all = @{lemma "(ALL x. P) == P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   536
  val elim_ex = @{lemma "(EX x. P) == P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   537
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   538
  fun elim_unused_conv ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   539
    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
   540
      (Conv.rewrs_conv [elim_all, elim_ex])))) ctxt
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   541
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   542
  fun elim_unused_tac ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   543
    REPEAT_ALL_NEW (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   544
      Tactic.match_tac [@{thm refl}, @{thm iff_allI}, @{thm iff_exI}]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   545
      ORELSE' CONVERSION (elim_unused_conv ctxt))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   546
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   547
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
   548
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   549
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   550
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   551
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   552
(* |- (ALL x1 ... xn. ~(x1 = t1 & ... xn = tn) | P x1 ... xn) = P t1 ... tn *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   553
fun dest_eq_res ctxt = Thm o try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   554
  named ctxt "fast" (T.by_tac (Classical.fast_tac HOL_cs))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   555
    (* FIXME: not very well tested *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   556
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   557
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   558
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   559
(* |- ~(ALL x1...xn. P x1...xn) | P a1...an *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   560
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   561
  val rule = @{lemma "~ P x | Q ==> ~(ALL x. P x) | Q" by fast}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   562
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   563
val quant_inst = Thm o T.by_tac (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   564
  REPEAT_ALL_NEW (Tactic.match_tac [rule])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   565
  THEN' Tactic.rtac @{thm excluded_middle})
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   566
end
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
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   570
(* c = SOME x. P x |- (EX x. P x) = P c
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   571
   c = SOME x. ~ P x |- ~(ALL x. P x) = ~ P c *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   572
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   573
  val elim_ex = @{lemma "EX x. P == P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   574
  val elim_all = @{lemma "~ (ALL x. P) == ~P" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   575
  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
   576
    by simp (intro eq_reflection some_eq_ex[symmetric])}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   577
  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
   578
    by (simp only: not_all) (intro eq_reflection some_eq_ex[symmetric])}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   579
  val sk_ex_rule = ((sk_ex, I), elim_ex)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   580
  and sk_all_rule = ((sk_all, Thm.dest_arg), elim_all)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   581
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   582
  fun dest f sk_rule = 
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   583
    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
   584
  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
   585
  fun pair2 (a, b) (c, d) = [(a, c), (b, d)]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   586
  fun inst_sk (sk_rule, f) p c =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   587
    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
   588
    |> (fn sk' => Thm.instantiate ([], (pair2 (dest f sk') (p, c))) sk')
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   589
    |> Conv.fconv_rule (Thm.beta_conversion true)
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 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
   592
    | 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
   593
        (sk_all_rule, Thm.dest_arg, L.negate)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   594
    | kind t = raise TERM ("skolemize", [t])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   595
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   596
  fun dest_abs_type (Abs (_, T, _)) = T
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   597
    | dest_abs_type t = raise TERM ("dest_abs_type", [t])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   598
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   599
  fun bodies_of thy lhs rhs =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   600
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   601
      val (rule, dest, make) = kind (Thm.term_of lhs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   602
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   603
      fun dest_body idx cbs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   604
        let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   605
          val cb = Thm.dest_arg (dest ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   606
          val T = dest_abs_type (Thm.term_of cb)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   607
          val cv = Thm.cterm_of thy (Var (("x", idx), T))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   608
          val cu = make (Drule.beta_conv cb cv)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   609
          val cbs' = (cv, cb) :: cbs
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   610
        in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   611
          (snd (Thm.first_order_match (cu, rhs)), rev cbs')
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   612
          handle Pattern.MATCH => dest_body (idx+1) cbs' cu
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   613
        end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   614
    in (rule, dest_body 1 [] lhs) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   615
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   616
  fun transitive f thm = Thm.transitive thm (f (Thm.rhs_of thm))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   617
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   618
  fun sk_step (rule, elim) (cv, mct, cb) ((is, thm), ctxt) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   619
    (case mct of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   620
      SOME ct =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   621
        ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   622
        |> T.make_hyp_def (inst_sk rule (Thm.instantiate_cterm ([], is) cb) ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   623
        |>> pair ((cv, ct) :: is) o Thm.transitive thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   624
    | NONE => ((is, transitive (Conv.rewr_conv elim) thm), ctxt))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   625
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   626
fun skolemize ct ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   627
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   628
    val (lhs, rhs) = Thm.dest_binop (Thm.dest_arg ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   629
    val (rule, (ctab, cbs)) = bodies_of (ProofContext.theory_of ctxt) lhs rhs
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   630
    fun lookup_var (cv, cb) = (cv, AList.lookup (op aconvc) ctab cv, cb)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   631
  in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   632
    (([], Thm.reflexive lhs), ctxt)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   633
    |> fold (sk_step rule) (map lookup_var cbs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   634
    |>> MetaEq o snd
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   635
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   636
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   637
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   638
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   639
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   640
(** theory proof rules **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   641
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   642
(* theory lemmas: linear arithmetic, arrays *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   643
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   644
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
   645
  T.by_abstraction (false, true) ctxt thms (fn ctxt' => T.by_tac (
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   646
    NAMED ctxt' "arith" (Arith_Data.arith_tac ctxt')
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   647
    ORELSE' NAMED ctxt' "simp+arith" (Simplifier.simp_tac simpset THEN_ALL_NEW
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   648
      Arith_Data.arith_tac ctxt')))]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   649
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   650
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   651
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   652
(* rewriting: prove equalities:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   653
     * ACI of conjunction/disjunction
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   654
     * contradiction, excluded middle
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   655
     * logical rewriting rules (for negation, implication, equivalence,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   656
         distinct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   657
     * normal forms for polynoms (integer/real arithmetic)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   658
     * quantifier elimination over linear arithmetic
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   659
     * ... ? **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   660
structure Z3_Simps = Named_Thms
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   661
(
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   662
  val name = "z3_simp"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   663
  val description = "simplification rules for Z3 proof reconstruction"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   664
)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   665
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   666
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   667
  fun spec_meta_eq_of thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   668
    (case try (fn th => th RS @{thm spec}) thm of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   669
      SOME thm' => spec_meta_eq_of thm'
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   670
    | NONE => mk_meta_eq thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   671
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   672
  fun prep (Thm thm) = spec_meta_eq_of thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   673
    | prep (MetaEq thm) = thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   674
    | prep (Literals (thm, _)) = spec_meta_eq_of thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   675
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   676
  fun unfold_conv ctxt ths =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   677
    Conv.arg_conv (Conv.binop_conv (T.unfold_eqs ctxt (map prep ths)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   678
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   679
  fun with_conv _ [] prv = prv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   680
    | with_conv ctxt ths prv = T.with_conv (unfold_conv ctxt ths) prv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   681
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   682
  val unfold_conv =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   683
    Conv.arg_conv (Conv.binop_conv (Conv.try_conv T.unfold_distinct_conv))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   684
  val prove_conj_disj_eq = T.with_conv unfold_conv L.prove_conj_disj_eq
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   685
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   686
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   687
fun rewrite ctxt simpset ths = Thm o with_conv ctxt ths (try_apply ctxt [] [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   688
  named ctxt "conj/disj/distinct" prove_conj_disj_eq,
37126
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   689
  T.by_abstraction (true, false) ctxt [] (fn ctxt' => T.by_tac (
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   690
    NAMED ctxt' "simp (logic)" (Simplifier.simp_tac simpset)
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   691
    THEN_ALL_NEW NAMED ctxt' "fast (logic)" (Classical.fast_tac HOL_cs))),
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   692
  T.by_abstraction (false, true) ctxt [] (fn ctxt' => T.by_tac (
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   693
    NAMED ctxt' "simp (theory)" (Simplifier.simp_tac simpset)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   694
    THEN_ALL_NEW (
37126
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   695
      NAMED ctxt' "fast (theory)" (Classical.fast_tac HOL_cs)
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   696
      ORELSE' NAMED ctxt' "arith (theory)" (Arith_Data.arith_tac ctxt')))),
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   697
  T.by_abstraction (true, true) ctxt [] (fn ctxt' => T.by_tac (
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   698
    NAMED ctxt' "simp (full)" (Simplifier.simp_tac simpset)
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   699
    THEN_ALL_NEW (
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   700
      NAMED ctxt' "fast (full)" (Classical.fast_tac HOL_cs)
fed6bbf35bac try logical and theory abstraction before full abstraction (avoids warnings of linarith)
boehmes
parents: 36936
diff changeset
   701
      ORELSE' NAMED ctxt' "arith (full)" (Arith_Data.arith_tac ctxt'))))])
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   702
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   703
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   704
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   705
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   706
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   707
(** proof reconstruction **)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   708
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   709
(* tracing and checking *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   710
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   711
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   712
  fun count_rules ptab =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   713
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   714
      fun count (_, Unproved _) (solved, total) = (solved, total + 1)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   715
        | count (_, Proved _) (solved, total) = (solved + 1, total + 1)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   716
    in Inttab.fold count ptab (0, 0) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   717
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   718
  fun header idx r (solved, total) = 
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   719
    "Z3: #" ^ string_of_int idx ^ ": " ^ P.string_of_rule r ^ " (goal " ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   720
    string_of_int (solved + 1) ^ " of " ^ string_of_int total ^ ")"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   721
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   722
  fun check ctxt idx r ps ct p =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   723
    let val thm = thm_of p |> tap (Thm.join_proofs o single)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   724
    in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   725
      if (Thm.cprop_of thm) aconvc ct then ()
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   726
      else z3_exn (Pretty.string_of (Pretty.big_list ("proof step failed: " ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   727
        quote (P.string_of_rule r) ^ " (#" ^ string_of_int idx ^ ")")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   728
          (pretty_goal ctxt (map (thm_of o fst) ps) (Thm.prop_of thm) @
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   729
           [Pretty.block [Pretty.str "expected: ",
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   730
            Syntax.pretty_term ctxt (Thm.term_of ct)]])))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   731
    end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   732
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   733
fun trace_rule idx prove r ps ct (cxp as (ctxt, ptab)) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   734
  let
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
   735
    val _ = SMT_Config.trace_msg ctxt (header idx r o count_rules) ptab
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   736
    val result as (p, (ctxt', _)) = prove r ps ct cxp
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
   737
    val _ = if not (Config.get ctxt' SMT_Config.trace) then ()
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   738
      else check ctxt' idx r ps ct p
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   739
  in result end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   740
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   741
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   742
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   743
(* overall reconstruction procedure *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   744
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
   745
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
   746
  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
   747
    quote (P.string_of_rule r))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   748
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
   749
  fun step assms simpset vars r ps ct (cxp as (cx, 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
   750
    (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
   751
      (* core 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
   752
      (P.TrueAxiom, _) => (Thm L.true_thm, 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
   753
    | (P.Asserted, _) => (asserted cx assms 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
   754
    | (P.Goal, _) => (asserted cx assms 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
   755
    | (P.ModusPonens, [(p, _), (q, _)]) => (mp q (thm_of 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
   756
    | (P.ModusPonensOeq, [(p, _), (q, _)]) => (mp q (thm_of 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
   757
    | (P.AndElim, [(p, i)]) => and_elim (p, i) ct ptab ||> pair cx
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
   758
    | (P.NotOrElim, [(p, i)]) => not_or_elim (p, i) ct ptab ||> pair cx
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
    | (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
   760
    | (P.Lemma, [(p, _)]) => (lemma (thm_of p) 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
   761
    | (P.UnitResolution, (p, _) :: ps) =>
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
        (unit_resolution (thm_of p) (map (thm_of o 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
   763
    | (P.IffTrue, [(p, _)]) => (iff_true (thm_of 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
   764
    | (P.IffFalse, [(p, _)]) => (iff_false (thm_of 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
   765
    | (P.Distributivity, _) => (distributivity cx 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
   766
    | (P.DefAxiom, _) => (def_axiom cx 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
   767
    | (P.IntroDef, _) => intro_def 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
   768
    | (P.ApplyDef, [(p, _)]) => (apply_def (thm_of 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
   769
    | (P.IffOeq, [(p, _)]) => (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
   770
    | (P.NnfPos, _) => (nnf cx vars (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
   771
    | (P.NnfNeg, _) => (nnf cx vars (map fst ps) ct, cxp)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   772
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
   773
      (* 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
   774
    | (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
   775
    | (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
   776
    | (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
   777
    | (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
   778
    | (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
   779
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
   780
      (* quantifier 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
   781
    | (P.QuantIntro, [(p, _)]) => (quant_intro vars p 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
   782
    | (P.PullQuant, _) => (pull_quant cx 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
   783
    | (P.PushQuant, _) => (push_quant cx 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
   784
    | (P.ElimUnusedVars, _) => (elim_unused_vars cx 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
   785
    | (P.DestEqRes, _) => (dest_eq_res cx 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
   786
    | (P.QuantInst, _) => (quant_inst 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
   787
    | (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
   788
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
   789
      (* theory rules *)
40516
516a367eb38c preliminary support for newer versions of Z3
boehmes
parents: 40424
diff changeset
   790
    | (P.ThLemma _, _) =>  (* 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
   791
        (th_lemma cx simpset (map (thm_of o 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
   792
    | (P.Rewrite, _) => (rewrite cx simpset [] 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
   793
    | (P.RewriteStar, ps) =>
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
   794
        (rewrite cx simpset (map fst ps) ct, cxp)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   795
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
   796
    | (P.NnfStar, _) => not_supported 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
   797
    | (P.CnfStar, _) => not_supported 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
    | (P.TransitivityStar, _) => not_supported 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
   799
    | (P.PullQuantStar, _) => not_supported r
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   800
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
   801
    | _ => 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
   802
       " has an unexpected number of arguments."))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   803
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
   804
  fun prove ctxt assms vars =
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
   805
    let
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
      val simpset = T.make_simpset ctxt (Z3_Simps.get ctxt)
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
   807
 
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
   808
      fun conclude idx rule prop (ps, 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
   809
        trace_rule idx (step assms simpset vars) rule ps prop 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
   810
        |-> (fn p => apsnd (Inttab.update (idx, Proved p)) #> pair p)
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
   811
 
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
   812
      fun lookup idx (cxp as (_, 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
   813
        (case Inttab.lookup ptab idx 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
   814
          SOME (Unproved (P.Proof_Step {rule, prems, prop})) =>
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
   815
            fold_map lookup prems 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
   816
            |>> map2 rpair prems
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
   817
            |> conclude idx rule prop
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
   818
        | SOME (Proved p) => (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
   819
        | NONE => z3_exn ("unknown proof id: " ^ quote (string_of_int idx)))
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
   820
 
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
   821
      fun result (p, (cx, _)) = (thm_of p, cx)
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
   822
    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
   823
      (fn idx => result o lookup idx o pair ctxt o Inttab.map (K Unproved))
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
   824
    end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   825
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
   826
  fun filter_assms ctxt assms 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
   827
    let
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
   828
      fun step r ct =
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
   829
        (case r 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
   830
          P.Asserted => insert (op =) (find_assm ctxt assms ct)
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
        | P.Goal => insert (op =) (find_assm ctxt assms ct)
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
   832
        | _ => I)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   833
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
   834
      fun lookup idx =
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
   835
        (case Inttab.lookup ptab idx 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
   836
          SOME (P.Proof_Step {rule, prems, prop}) =>
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
   837
            fold lookup prems #> step rule prop
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
   838
        | NONE => z3_exn ("unknown proof id: " ^ quote (string_of_int idx)))
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
   839
    in lookup end
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
   840
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
   841
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
   842
fun reconstruct ctxt {typs, terms, unfolds, assms} output =
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
   843
  let
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
   844
    val (idx, (ptab, vars, cx)) = P.parse ctxt typs terms output
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
   845
    val assms' = prepare_assms cx unfolds assms
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   846
  in
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
   847
    if Config.get cx SMT_Config.filter_only_facts
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
   848
    then ((filter_assms cx assms' ptab idx [], @{thm TrueI}), cx)
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
   849
    else apfst (pair []) (prove cx assms' vars idx ptab)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   850
  end
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
end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   853
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
   854
val setup = z3_rules_setup #> Z3_Simps.setup
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   855
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   856
end