src/HOL/Tools/SMT/z3_interface.ML
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 74817 1fd8705503b4
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
     1
(*  Title:      HOL/Tools/SMT/z3_interface.ML
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     3
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     4
Interface to Z3 based on a relaxed version of SMT-LIB.
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     5
*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     6
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
     7
signature Z3_INTERFACE =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     8
sig
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
     9
  val smtlib_z3C: SMT_Util.class
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    10
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    11
  datatype sym = Sym of string * sym list
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    12
  type mk_builtins = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    13
    mk_builtin_typ: sym -> typ option,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    14
    mk_builtin_num: theory -> int -> typ -> cterm option,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
    mk_builtin_fun: theory -> sym -> cterm list -> cterm option }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    16
  val add_mk_builtins: mk_builtins -> Context.generic -> Context.generic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
  val mk_builtin_typ: Proof.context -> sym -> typ option
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    18
  val mk_builtin_num: Proof.context -> int -> typ -> cterm option
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
  val mk_builtin_fun: Proof.context -> sym -> cterm list -> cterm option
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    20
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    21
  val is_builtin_theory_term: Proof.context -> term -> bool
57229
blanchet
parents: 56103
diff changeset
    22
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    23
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
    24
structure Z3_Interface: Z3_INTERFACE =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    26
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 63950
diff changeset
    27
val z3C = ["z3"]
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 63950
diff changeset
    28
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    29
val smtlib_z3C = SMTLIB_Interface.smtlibC @ SMTLIB_Interface.bvsmlibC @ z3C
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    30
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    32
(* interface *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    34
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
  fun translate_config ctxt =
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 63950
diff changeset
    36
    {order = SMT_Util.First_Order,
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    37
     logic = K (K ""),
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 63950
diff changeset
    38
     fp_kinds = [BNF_Util.Least_FP],
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 63950
diff changeset
    39
     serialize = #serialize (SMTLIB_Interface.translate_config SMT_Util.First_Order ctxt)}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
    41
  fun is_div_mod \<^Const_>\<open>divide \<^Type>\<open>int\<close>\<close> = true
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
    42
    | is_div_mod \<^Const>\<open>modulo \<^Type>\<open>int\<close>\<close> = true
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
    | is_div_mod _ = false
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    44
56103
6689512f3710 move lemmas to theory file, towards textual proof reconstruction
blanchet
parents: 56101
diff changeset
    45
  val have_int_div_mod = exists (Term.exists_subterm is_div_mod o Thm.prop_of)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    46
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    47
  fun add_div_mod _ (thms, extra_thms) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    48
    if have_int_div_mod thms orelse have_int_div_mod extra_thms then
56101
blanchet
parents: 56090
diff changeset
    49
      (thms, @{thms div_as_z3div mod_as_z3mod} @ extra_thms)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
    else (thms, extra_thms)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
  val setup_builtins =
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
    53
    SMT_Builtin.add_builtin_fun' smtlib_z3C (\<^Const>\<open>times \<^Type>\<open>int\<close>\<close>, "*") #>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
    54
    SMT_Builtin.add_builtin_fun' smtlib_z3C (\<^Const>\<open>z3div\<close>, "div") #>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
    55
    SMT_Builtin.add_builtin_fun' smtlib_z3C (\<^Const>\<open>z3mod\<close>, "mod")
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    56
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
val _ = Theory.setup (Context.theory_map (
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
  setup_builtins #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
    60
  SMT_Normalize.add_extra_norm (smtlib_z3C, add_div_mod) #>
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
    61
  SMT_Translate.add_config (smtlib_z3C, translate_config)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    63
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    64
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
(* constructors *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    67
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    68
datatype sym = Sym of string * sym list
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    69
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    70
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    71
(** additional constructors **)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    72
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
type mk_builtins = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    74
  mk_builtin_typ: sym -> typ option,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    75
  mk_builtin_num: theory -> int -> typ -> cterm option,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    76
  mk_builtin_fun: theory -> sym -> cterm list -> cterm option }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    77
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    78
fun chained _ [] = NONE
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    79
  | chained f (b :: bs) = (case f b of SOME y => SOME y | NONE => chained f bs)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    80
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    81
fun chained_mk_builtin_typ bs sym =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    82
  chained (fn {mk_builtin_typ=mk, ...} : mk_builtins => mk sym) bs
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    83
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    84
fun chained_mk_builtin_num ctxt bs i T =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    85
  let val thy = Proof_Context.theory_of ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    86
  in chained (fn {mk_builtin_num=mk, ...} : mk_builtins => mk thy i T) bs end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    87
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
fun chained_mk_builtin_fun ctxt bs s cts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    89
  let val thy = Proof_Context.theory_of ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    90
  in chained (fn {mk_builtin_fun=mk, ...} : mk_builtins => mk thy s cts) bs end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    91
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    92
fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    93
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    94
structure Mk_Builtins = Generic_Data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    95
(
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    96
  type T = (int * mk_builtins) list
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    97
  val empty = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    98
  fun merge data = Ord_List.merge fst_int_ord data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    99
)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   100
56103
6689512f3710 move lemmas to theory file, towards textual proof reconstruction
blanchet
parents: 56101
diff changeset
   101
fun add_mk_builtins mk = Mk_Builtins.map (Ord_List.insert fst_int_ord (serial (), mk))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   102
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
fun get_mk_builtins ctxt = map snd (Mk_Builtins.get (Context.Proof ctxt))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   104
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
(** basic and additional constructors **)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   108
fun mk_builtin_typ _ (Sym ("Bool", _)) = SOME \<^typ>\<open>bool\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   109
  | mk_builtin_typ _ (Sym ("Int", _)) = SOME \<^typ>\<open>int\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   110
  | mk_builtin_typ _ (Sym ("bool", _)) = SOME \<^typ>\<open>bool\<close>  (*FIXME: legacy*)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   111
  | mk_builtin_typ _ (Sym ("int", _)) = SOME \<^typ>\<open>int\<close>  (*FIXME: legacy*)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   112
  | mk_builtin_typ ctxt sym = chained_mk_builtin_typ (get_mk_builtins ctxt) sym
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   113
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   114
fun mk_builtin_num _ i \<^typ>\<open>int\<close> = SOME (Numeral.mk_cnumber \<^ctyp>\<open>int\<close> i)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   115
  | mk_builtin_num ctxt i T =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   116
      chained_mk_builtin_num ctxt (get_mk_builtins ctxt) i T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   117
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   118
val mk_true = \<^cterm>\<open>\<not> False\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   119
val mk_false = \<^cterm>\<open>False\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   120
val mk_not = Thm.apply \<^cterm>\<open>HOL.Not\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   121
val mk_implies = Thm.mk_binop \<^cterm>\<open>HOL.implies\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   122
val mk_iff = Thm.mk_binop \<^cterm>\<open>(=) :: bool \<Rightarrow> _\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   123
val conj = \<^cterm>\<open>HOL.conj\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   124
val disj = \<^cterm>\<open>HOL.disj\<close>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   125
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   126
fun mk_nary _ cu [] = cu
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   127
  | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   128
70159
57503fe1b0ff tuned signature;
wenzelm
parents: 70150
diff changeset
   129
val eq = SMT_Util.mk_const_pat \<^theory> \<^const_name>\<open>HOL.eq\<close> Thm.dest_ctyp0
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
   130
fun mk_eq ct cu = Thm.mk_binop (SMT_Util.instT' ct eq) ct cu
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   131
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   132
val if_term =
70159
57503fe1b0ff tuned signature;
wenzelm
parents: 70150
diff changeset
   133
  SMT_Util.mk_const_pat \<^theory> \<^const_name>\<open>If\<close> (Thm.dest_ctyp0 o Thm.dest_ctyp1)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
   134
fun mk_if cc ct = Thm.mk_binop (Thm.apply (SMT_Util.instT' ct if_term) cc) ct
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   135
70159
57503fe1b0ff tuned signature;
wenzelm
parents: 70150
diff changeset
   136
val access = SMT_Util.mk_const_pat \<^theory> \<^const_name>\<open>fun_app\<close> Thm.dest_ctyp0
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
   137
fun mk_access array = Thm.apply (SMT_Util.instT' array access) array
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   138
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56078
diff changeset
   139
val update =
70159
57503fe1b0ff tuned signature;
wenzelm
parents: 70150
diff changeset
   140
  SMT_Util.mk_const_pat \<^theory> \<^const_name>\<open>fun_upd\<close> (Thm.dest_ctyp o Thm.dest_ctyp0)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
fun mk_update array index value =
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 58360
diff changeset
   142
  let val cTs = Thm.dest_ctyp (Thm.ctyp_of_cterm array)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
   143
  in Thm.apply (Thm.mk_binop (SMT_Util.instTs cTs update) array index) value end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   144
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   145
val mk_uminus = Thm.apply \<^cterm>\<open>uminus :: int \<Rightarrow> _\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   146
val add = \<^cterm>\<open>(+) :: int \<Rightarrow> _\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   147
val int0 = Numeral.mk_cnumber \<^ctyp>\<open>int\<close> 0
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   148
val mk_sub = Thm.mk_binop \<^cterm>\<open>(-) :: int \<Rightarrow> _\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   149
val mk_mul = Thm.mk_binop \<^cterm>\<open>(*) :: int \<Rightarrow> _\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   150
val mk_div = Thm.mk_binop \<^cterm>\<open>z3div\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   151
val mk_mod = Thm.mk_binop \<^cterm>\<open>z3mod\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   152
val mk_lt = Thm.mk_binop \<^cterm>\<open>(<) :: int \<Rightarrow> _\<close>
8d0294d877bd clarified antiquotations;
wenzelm
parents: 70159
diff changeset
   153
val mk_le = Thm.mk_binop \<^cterm>\<open>(\<le>) :: int \<Rightarrow> _\<close>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   154
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   155
fun mk_builtin_fun ctxt sym cts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   156
  (case (sym, cts) of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   157
    (Sym ("true", _), []) => SOME mk_true
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   158
  | (Sym ("false", _), []) => SOME mk_false
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   159
  | (Sym ("not", _), [ct]) => SOME (mk_not ct)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   160
  | (Sym ("and", _), _) => SOME (mk_nary conj mk_true cts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   161
  | (Sym ("or", _), _) => SOME (mk_nary disj mk_false cts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   162
  | (Sym ("implies", _), [ct, cu]) => SOME (mk_implies ct cu)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
  | (Sym ("iff", _), [ct, cu]) => SOME (mk_iff ct cu)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   164
  | (Sym ("~", _), [ct, cu]) => SOME (mk_iff ct cu)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
  | (Sym ("xor", _), [ct, cu]) => SOME (mk_not (mk_iff ct cu))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
  | (Sym ("if", _), [ct1, ct2, ct3]) => SOME (mk_if ct1 ct2 ct3)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   167
  | (Sym ("ite", _), [ct1, ct2, ct3]) => SOME (mk_if ct1 ct2 ct3) (* FIXME: remove *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   168
  | (Sym ("=", _), [ct, cu]) => SOME (mk_eq ct cu)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   169
  | (Sym ("select", _), [ca, ck]) => SOME (Thm.apply (mk_access ca) ck)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   170
  | (Sym ("store", _), [ca, ck, cv]) => SOME (mk_update ca ck cv)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   171
  | _ =>
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 58360
diff changeset
   172
    (case (sym, try (Thm.typ_of_cterm o hd) cts, cts) of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   173
      (Sym ("+", _), SOME \<^typ>\<open>int\<close>, _) => SOME (mk_nary add int0 cts)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   174
    | (Sym ("-", _), SOME \<^typ>\<open>int\<close>, [ct]) => SOME (mk_uminus ct)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   175
    | (Sym ("-", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_sub ct cu)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   176
    | (Sym ("*", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_mul ct cu)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   177
    | (Sym ("div", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_div ct cu)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   178
    | (Sym ("mod", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_mod ct cu)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   179
    | (Sym ("<", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_lt ct cu)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   180
    | (Sym ("<=", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_le ct cu)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   181
    | (Sym (">", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_lt cu ct)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66551
diff changeset
   182
    | (Sym (">=", _), SOME \<^typ>\<open>int\<close>, [ct, cu]) => SOME (mk_le cu ct)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   183
    | _ => chained_mk_builtin_fun ctxt (get_mk_builtins ctxt) sym cts))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   184
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   185
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   186
(* abstraction *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   187
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   188
fun is_builtin_theory_term ctxt t =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
   189
  if SMT_Builtin.is_builtin_num ctxt t then true
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   190
  else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   191
    (case Term.strip_comb t of
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57238
diff changeset
   192
      (Const c, ts) => SMT_Builtin.is_builtin_fun ctxt c ts
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   193
    | _ => false)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   194
57229
blanchet
parents: 56103
diff changeset
   195
end;