src/HOL/Tools/SMT/z3_proof_tools.ML
author wenzelm
Sun, 07 Nov 2010 23:32:26 +0100
changeset 40405 42671298f037
parent 40274 6486c610a549
child 40579 98ebd2300823
permissions -rw-r--r--
tweaked pdf setup to allow modification of \pdfminorversion;
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_tools.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
Helper functions required for Z3 proof reconstruction.
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_TOOLS =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     8
sig
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     9
  (* accessing and modifying terms *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    10
  val term_of: cterm -> term
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    11
  val prop_of: thm -> term
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    12
  val mk_prop: cterm -> cterm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    13
  val as_meta_eq: cterm -> cterm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    14
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    15
  (* theorem nets *)
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    16
  val thm_net_of: ('a -> thm) -> 'a list -> 'a Net.net
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    17
  val net_instance': ((thm -> thm option) -> 'a -> 'a option) -> 'a Net.net ->
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    18
    cterm -> 'a option
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    19
  val net_instance: thm Net.net -> cterm -> thm option
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    20
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    21
  (* proof combinators *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    22
  val under_assumption: (thm -> thm) -> cterm -> thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    23
  val with_conv: conv -> (cterm -> thm) -> cterm -> thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    24
  val discharge: thm -> thm -> thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    25
  val varify: string list -> thm -> thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    26
  val unfold_eqs: Proof.context -> thm list -> conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    27
  val match_instantiate: (cterm -> cterm) -> cterm -> thm -> thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    28
  val by_tac: (int -> tactic) -> cterm -> thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    29
  val make_hyp_def: thm -> Proof.context -> thm * Proof.context
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    30
  val by_abstraction: bool * bool -> Proof.context -> thm list ->
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    31
    (Proof.context -> cterm -> thm) -> cterm -> thm
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    32
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    33
  (* a faster COMP *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    34
  type compose_data
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    35
  val precompose: (cterm -> cterm list) -> thm -> compose_data
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    36
  val precompose2: (cterm -> cterm * cterm) -> thm -> compose_data
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    37
  val compose: compose_data -> thm -> thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    38
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    39
  (* unfolding of 'distinct' *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    40
  val unfold_distinct_conv: conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    41
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    42
  (* simpset *)
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    43
  val add_simproc: Simplifier.simproc -> Context.generic -> Context.generic
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
  val make_simpset: Proof.context -> thm list -> simpset
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    45
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    46
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    47
structure Z3_Proof_Tools: Z3_PROOF_TOOLS =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    48
struct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    50
structure I = Z3_Interface
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    51
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    52
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    53
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    54
(* accessing terms *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    55
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    56
val dest_prop = (fn @{term Trueprop} $ t => t | t => t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    57
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    58
fun term_of ct = dest_prop (Thm.term_of ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    59
fun prop_of thm = dest_prop (Thm.prop_of thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    60
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    61
val mk_prop = Thm.capply @{cterm Trueprop}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    62
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    63
val eq = I.mk_inst_pair I.destT1 @{cpat "op =="}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    64
fun mk_meta_eq_cterm ct cu = Thm.mk_binop (I.instT' ct eq) ct cu
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    65
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    66
fun as_meta_eq ct = uncurry mk_meta_eq_cterm (Thm.dest_binop (Thm.dest_arg ct))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    67
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    68
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    69
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    70
(* theorem nets *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    71
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    72
fun thm_net_of f xthms =
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    73
  let fun insert xthm = Net.insert_term (K false) (Thm.prop_of (f xthm), xthm)
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    74
  in fold insert xthms Net.empty end
36898
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 maybe_instantiate ct thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    77
  try Thm.first_order_match (Thm.cprop_of thm, ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    78
  |> Option.map (fn inst => Thm.instantiate inst thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    79
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    80
fun net_instance' f net ct =
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    81
  let fun first_of f xthms ct = get_first (f (maybe_instantiate ct)) xthms 
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    82
  in first_of f (Net.match_term net (Thm.term_of ct)) ct end
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 38864
diff changeset
    83
val net_instance = net_instance' I
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    84
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    85
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    86
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    87
(* proof combinators *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    88
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    89
fun under_assumption f ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    90
  let val ct' = mk_prop ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    91
  in Thm.implies_intr ct' (f (Thm.assume ct')) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    92
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    93
fun with_conv conv prove ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    94
  let val eq = Thm.symmetric (conv ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    95
  in Thm.equal_elim eq (prove (Thm.lhs_of eq)) end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    96
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    97
fun discharge p pq = Thm.implies_elim pq p
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    98
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    99
fun varify vars = Drule.generalize ([], vars)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   100
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   101
fun unfold_eqs _ [] = Conv.all_conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   102
  | unfold_eqs ctxt eqs =
36936
c52d1c130898 incorporated further conversions and conversionals, after some minor tuning;
wenzelm
parents: 36899
diff changeset
   103
      Conv.top_sweep_conv (K (Conv.rewrs_conv eqs)) ctxt
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   104
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   105
fun match_instantiate f ct thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   106
  Thm.instantiate (Thm.match (f (Thm.cprop_of thm), ct)) thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   107
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   108
fun by_tac tac ct = Goal.norm_result (Goal.prove_internal [] ct (K (tac 1)))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   109
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   110
(* |- c x == t x ==> P (c x)  ~~>  c == t |- P (c x) *) 
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   111
fun make_hyp_def thm ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   112
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   113
    val (lhs, rhs) = Thm.dest_binop (Thm.cprem_of thm 1)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   114
    val (cf, cvs) = Drule.strip_comb lhs
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   115
    val eq = mk_meta_eq_cterm cf (fold_rev Thm.cabs cvs rhs)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   116
    fun apply cv th =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   117
      Thm.combination th (Thm.reflexive cv)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   118
      |> Conv.fconv_rule (Conv.arg_conv (Thm.beta_conversion false))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   119
  in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   120
    yield_singleton Assumption.add_assumes eq ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   121
    |>> Thm.implies_elim thm o fold apply cvs
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   122
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   123
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   124
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   125
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   126
(* abstraction *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   127
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   128
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   129
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   130
fun typ_of ct = #T (Thm.rep_cterm ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   131
fun certify ctxt = Thm.cterm_of (ProofContext.theory_of ctxt)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   132
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   133
fun abs_context ctxt = (ctxt, Termtab.empty, 1, false)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   134
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   135
fun context_of (ctxt, _, _, _) = ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   136
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   137
fun replace (_, (cv, ct)) = Thm.forall_elim ct o Thm.forall_intr cv
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   138
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   139
fun abs_instantiate (_, tab, _, beta_norm) =
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   140
  fold replace (Termtab.dest tab) #>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   141
  beta_norm ? Conv.fconv_rule (Thm.beta_conversion true)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   142
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   143
fun lambda_abstract cvs t =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   144
  let
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   145
    val frees = map Free (Term.add_frees t [])
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   146
    val cvs' = filter (fn cv => member (op aconv) frees (Thm.term_of cv)) cvs
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   147
    val vs = map (Term.dest_Free o Thm.term_of) cvs'
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   148
  in (Term.list_abs_free (vs, t), cvs') end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   149
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   150
fun fresh_abstraction cvs ct (cx as (ctxt, tab, idx, beta_norm)) =
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   151
  let val (t, cvs') = lambda_abstract cvs (Thm.term_of ct)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   152
  in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   153
    (case Termtab.lookup tab t of
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   154
      SOME (cv, _) => (Drule.list_comb (cv, cvs'), cx)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   155
    | NONE =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   156
        let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   157
          val (n, ctxt') = yield_singleton Variable.variant_fixes "x" ctxt
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   158
          val cv = certify ctxt' (Free (n, map typ_of cvs' ---> typ_of ct))
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   159
          val cu = Drule.list_comb (cv, cvs')
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   160
          val e = (t, (cv, fold_rev Thm.cabs cvs' ct))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   161
          val beta_norm' = beta_norm orelse not (null cvs')
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   162
        in (cu, (ctxt', Termtab.update e tab, idx + 1, beta_norm')) end)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   163
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   164
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   165
fun abs_comb f g cvs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   166
  let val (cf, cu) = Thm.dest_comb ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   167
  in f cvs cf ##>> g cvs cu #>> uncurry Thm.capply end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   168
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   169
fun abs_arg f = abs_comb (K pair) f
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   170
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   171
fun abs_args f cvs ct =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   172
  (case Thm.term_of ct of
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   173
    _ $ _ => abs_comb (abs_args f) f cvs ct
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   174
  | _ => pair ct)
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   175
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   176
fun abs_list f g cvs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   177
  (case Thm.term_of ct of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   178
    Const (@{const_name Nil}, _) => pair ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   179
  | Const (@{const_name Cons}, _) $ _ $ _ =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   180
      abs_comb (abs_arg f) (abs_list f g) cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   181
  | _ => g cvs ct)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   182
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   183
fun abs_abs f cvs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   184
  let val (cv, cu) = Thm.dest_abs NONE ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   185
  in f (cv :: cvs) cu #>> Thm.cabs cv end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   186
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   187
val is_atomic = (fn _ $ _ => false | Abs _ => false | _ => true)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   188
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   189
fun abstract (ext_logic, with_theories) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   190
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   191
    fun abstr1 cvs ct = abs_arg abstr cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   192
    and abstr2 cvs ct = abs_comb abstr1 abstr cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   193
    and abstr3 cvs ct = abs_comb abstr2 abstr cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   194
    and abstr_abs cvs ct = abs_arg (abs_abs abstr) cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   195
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   196
    and abstr cvs ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   197
      (case Thm.term_of ct of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   198
        @{term Trueprop} $ _ => abstr1 cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   199
      | @{term "op ==>"} $ _ $ _ => abstr2 cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   200
      | @{term True} => pair ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   201
      | @{term False} => pair ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   202
      | @{term Not} $ _ => abstr1 cvs ct
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   203
      | @{term HOL.conj} $ _ $ _ => abstr2 cvs ct
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   204
      | @{term HOL.disj} $ _ $ _ => abstr2 cvs ct
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38715
diff changeset
   205
      | @{term HOL.implies} $ _ $ _ => abstr2 cvs ct
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   206
      | Const (@{const_name HOL.eq}, _) $ _ $ _ => abstr2 cvs ct
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
   207
      | Const (@{const_name SMT.distinct}, _) $ _ =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   208
          if ext_logic then abs_arg (abs_list abstr fresh_abstraction) cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   209
          else fresh_abstraction cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   210
      | Const (@{const_name If}, _) $ _ $ _ $ _ =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   211
          if ext_logic then abstr3 cvs ct else fresh_abstraction cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   212
      | Const (@{const_name All}, _) $ _ =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   213
          if ext_logic then abstr_abs cvs ct else fresh_abstraction cvs ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   214
      | Const (@{const_name Ex}, _) $ _ =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   215
          if ext_logic then abstr_abs cvs ct else fresh_abstraction cvs ct
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   216
      | t => (fn cx =>
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   217
          if is_atomic t orelse can HOLogic.dest_number t then (ct, cx)
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   218
          else if with_theories andalso
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   219
            I.is_builtin_theory_term (context_of cx) t
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   220
          then abs_args abstr cvs ct cx
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   221
          else fresh_abstraction cvs ct cx))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   222
  in abstr [] end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   223
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   224
fun with_prems thms f ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   225
  fold_rev (Thm.mk_binop @{cterm "op ==>"} o Thm.cprop_of) thms ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   226
  |> f
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   227
  |> fold (fn prem => fn th => Thm.implies_elim th prem) thms
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   228
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   229
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   230
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   231
fun by_abstraction mode ctxt thms prove = with_prems thms (fn ct =>
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   232
  let val (cu, cx) = abstract mode ct (abs_context ctxt)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   233
  in abs_instantiate cx (prove (context_of cx) cu) end)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   234
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   235
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   236
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   237
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   238
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   239
(* a faster COMP *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   240
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   241
type compose_data = cterm list * (cterm -> cterm list) * thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   242
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   243
fun list2 (x, y) = [x, y]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   244
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   245
fun precompose f rule = (f (Thm.cprem_of rule 1), f, rule)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   246
fun precompose2 f rule = precompose (list2 o f) rule
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   247
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   248
fun compose (cvs, f, rule) thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   249
  discharge thm (Thm.instantiate ([], cvs ~~ f (Thm.cprop_of thm)) rule)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   250
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   251
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   252
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   253
(* unfolding of 'distinct' *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   254
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   255
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   256
  val set1 = @{lemma "x ~: set [] == ~False" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   257
  val set2 = @{lemma "x ~: set [x] == False" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   258
  val set3 = @{lemma "x ~: set [y] == x ~= y" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   259
  val set4 = @{lemma "x ~: set (x # ys) == False" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   260
  val set5 = @{lemma "x ~: set (y # ys) == x ~= y & x ~: set ys" by simp}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   261
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   262
  fun set_conv ct =
36936
c52d1c130898 incorporated further conversions and conversionals, after some minor tuning;
wenzelm
parents: 36899
diff changeset
   263
    (Conv.rewrs_conv [set1, set2, set3, set4] else_conv
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   264
    (Conv.rewr_conv set5 then_conv Conv.arg_conv set_conv)) ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   265
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
   266
  val dist1 = @{lemma "SMT.distinct [] == ~False" by (simp add: distinct_def)}
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
   267
  val dist2 = @{lemma "SMT.distinct [x] == ~False" by (simp add: distinct_def)}
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
   268
  val dist3 = @{lemma "SMT.distinct (x # xs) == x ~: set xs & distinct xs"
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
   269
    by (simp add: distinct_def)}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   270
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   271
  fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   272
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   273
fun unfold_distinct_conv ct =
36936
c52d1c130898 incorporated further conversions and conversionals, after some minor tuning;
wenzelm
parents: 36899
diff changeset
   274
  (Conv.rewrs_conv [dist1, dist2] else_conv
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   275
  (Conv.rewr_conv dist3 then_conv binop_conv set_conv unfold_distinct_conv)) ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   276
end
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
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   280
(* simpset *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   281
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   282
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   283
  val antisym_le1 = mk_meta_eq @{thm order_class.antisym_conv}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   284
  val antisym_le2 = mk_meta_eq @{thm linorder_class.antisym_conv2}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   285
  val antisym_less1 = mk_meta_eq @{thm linorder_class.antisym_conv1}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   286
  val antisym_less2 = mk_meta_eq @{thm linorder_class.antisym_conv3}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   287
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   288
  fun eq_prop t thm = HOLogic.mk_Trueprop t aconv Thm.prop_of thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   289
  fun dest_binop ((c as Const _) $ t $ u) = (c, t, u)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   290
    | dest_binop t = raise TERM ("dest_binop", [t])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   291
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   292
  fun prove_antisym_le ss t =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   293
    let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   294
      val (le, r, s) = dest_binop t
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   295
      val less = Const (@{const_name less}, Term.fastype_of le)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   296
      val prems = Simplifier.prems_of_ss ss
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   297
    in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   298
      (case find_first (eq_prop (le $ s $ r)) prems of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   299
        NONE =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   300
          find_first (eq_prop (HOLogic.mk_not (less $ r $ s))) prems
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   301
          |> Option.map (fn thm => thm RS antisym_less1)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   302
      | SOME thm => SOME (thm RS antisym_le1))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   303
    end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   304
    handle THM _ => NONE
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   305
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   306
  fun prove_antisym_less ss t =
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
      val (less, r, s) = dest_binop (HOLogic.dest_not t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   309
      val le = Const (@{const_name less_eq}, Term.fastype_of less)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   310
      val prems = prems_of_ss ss
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   311
    in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   312
      (case find_first (eq_prop (le $ r $ s)) prems of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   313
        NONE =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   314
          find_first (eq_prop (HOLogic.mk_not (less $ s $ r))) prems
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   315
          |> Option.map (fn thm => thm RS antisym_less2)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   316
      | SOME thm => SOME (thm RS antisym_le2))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   317
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   318
  handle THM _ => NONE
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   319
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   320
  val basic_simpset = HOL_ss addsimps @{thms field_simps}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   321
    addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}]
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   322
    addsimps @{thms arith_special} addsimps @{thms less_bin_simps}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   323
    addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   324
    addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   325
    addsimps @{thms minus_bin_simps} addsimps @{thms pred_bin_simps}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   326
    addsimps @{thms mult_bin_simps} addsimps @{thms iszero_simps}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   327
    addsimps @{thms array_rules}
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 36936
diff changeset
   328
    addsimps @{thms z3div_def} addsimps @{thms z3mod_def}
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   329
    addsimprocs [
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 37151
diff changeset
   330
      Simplifier.simproc_global @{theory} "fast_int_arith" [
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   331
        "(m::int) < n", "(m::int) <= n", "(m::int) = n"] (K Lin_Arith.simproc),
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 37151
diff changeset
   332
      Simplifier.simproc_global @{theory} "antisym_le" ["(x::'a::order) <= y"]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   333
        (K prove_antisym_le),
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 37151
diff changeset
   334
      Simplifier.simproc_global @{theory} "antisym_less" ["~ (x::'a::linorder) < y"]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   335
        (K prove_antisym_less)]
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   336
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   337
  structure Simpset = Generic_Data
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   338
  (
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   339
    type T = simpset
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   340
    val empty = basic_simpset
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   341
    val extend = I
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   342
    val merge = Simplifier.merge_ss
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   343
  )
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   344
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   345
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   346
fun add_simproc simproc = Simpset.map (fn ss => ss addsimprocs [simproc])
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   347
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   348
fun make_simpset ctxt rules =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   349
  Simplifier.context ctxt (Simpset.get (Context.Proof ctxt)) addsimps rules
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   350
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   351
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   352
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   353
end