src/HOL/Library/Old_SMT/smt_real.ML
changeset 58057 883f3c4c928e
parent 58055 625bdd5c70b2
equal deleted inserted replaced
58056:fc6dd578d506 58057:883f3c4c928e
       
     1 (*  Title:      HOL/Library/Old_SMT/smt_real.ML
       
     2     Author:     Sascha Boehme, TU Muenchen
       
     3 
       
     4 SMT setup for reals.
       
     5 *)
       
     6 
       
     7 signature SMT_REAL =
       
     8 sig
       
     9   val setup: theory -> theory
       
    10 end
       
    11 
       
    12 structure SMT_Real: SMT_REAL =
       
    13 struct
       
    14 
       
    15 
       
    16 (* SMT-LIB logic *)
       
    17 
       
    18 fun smtlib_logic ts =
       
    19   if exists (Term.exists_type (Term.exists_subtype (equal @{typ real}))) ts
       
    20   then SOME "AUFLIRA"
       
    21   else NONE
       
    22 
       
    23 
       
    24 (* SMT-LIB and Z3 built-ins *)
       
    25 
       
    26 local
       
    27   fun real_num _ i = SOME (string_of_int i ^ ".0")
       
    28 
       
    29   fun is_linear [t] = SMT_Utils.is_number t
       
    30     | is_linear [t, u] = SMT_Utils.is_number t orelse SMT_Utils.is_number u
       
    31     | is_linear _ = false
       
    32 
       
    33   fun mk_times ts = Term.list_comb (@{const times (real)}, ts)
       
    34 
       
    35   fun times _ _ ts = if is_linear ts then SOME ("*", 2, ts, mk_times) else NONE
       
    36 in
       
    37 
       
    38 val setup_builtins =
       
    39   SMT_Builtin.add_builtin_typ SMTLIB_Interface.smtlibC
       
    40     (@{typ real}, K (SOME "Real"), real_num) #>
       
    41   fold (SMT_Builtin.add_builtin_fun' SMTLIB_Interface.smtlibC) [
       
    42     (@{const less (real)}, "<"),
       
    43     (@{const less_eq (real)}, "<="),
       
    44     (@{const uminus (real)}, "~"),
       
    45     (@{const plus (real)}, "+"),
       
    46     (@{const minus (real)}, "-") ] #>
       
    47   SMT_Builtin.add_builtin_fun SMTLIB_Interface.smtlibC
       
    48     (Term.dest_Const @{const times (real)}, times) #>
       
    49   SMT_Builtin.add_builtin_fun' Z3_Interface.smtlib_z3C
       
    50     (@{const times (real)}, "*") #>
       
    51   SMT_Builtin.add_builtin_fun' Z3_Interface.smtlib_z3C
       
    52     (@{const divide (real)}, "/")
       
    53 
       
    54 end
       
    55 
       
    56 
       
    57 (* Z3 constructors *)
       
    58 
       
    59 local
       
    60   fun z3_mk_builtin_typ (Z3_Interface.Sym ("Real", _)) = SOME @{typ real}
       
    61     | z3_mk_builtin_typ (Z3_Interface.Sym ("real", _)) = SOME @{typ real}
       
    62         (*FIXME: delete*)
       
    63     | z3_mk_builtin_typ _ = NONE
       
    64 
       
    65   fun z3_mk_builtin_num _ i T =
       
    66     if T = @{typ real} then SOME (Numeral.mk_cnumber @{ctyp real} i)
       
    67     else NONE
       
    68 
       
    69   fun mk_nary _ cu [] = cu
       
    70     | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts)
       
    71 
       
    72   val mk_uminus = Thm.apply (Thm.cterm_of @{theory} @{const uminus (real)})
       
    73   val add = Thm.cterm_of @{theory} @{const plus (real)}
       
    74   val real0 = Numeral.mk_cnumber @{ctyp real} 0
       
    75   val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (real)})
       
    76   val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (real)})
       
    77   val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const divide (real)})
       
    78   val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (real)})
       
    79   val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (real)})
       
    80 
       
    81   fun z3_mk_builtin_fun (Z3_Interface.Sym ("-", _)) [ct] = SOME (mk_uminus ct)
       
    82     | z3_mk_builtin_fun (Z3_Interface.Sym ("+", _)) cts =
       
    83         SOME (mk_nary add real0 cts)
       
    84     | z3_mk_builtin_fun (Z3_Interface.Sym ("-", _)) [ct, cu] =
       
    85         SOME (mk_sub ct cu)
       
    86     | z3_mk_builtin_fun (Z3_Interface.Sym ("*", _)) [ct, cu] =
       
    87         SOME (mk_mul ct cu)
       
    88     | z3_mk_builtin_fun (Z3_Interface.Sym ("/", _)) [ct, cu] =
       
    89         SOME (mk_div ct cu)
       
    90     | z3_mk_builtin_fun (Z3_Interface.Sym ("<", _)) [ct, cu] =
       
    91         SOME (mk_lt ct cu)
       
    92     | z3_mk_builtin_fun (Z3_Interface.Sym ("<=", _)) [ct, cu] =
       
    93         SOME (mk_le ct cu)
       
    94     | z3_mk_builtin_fun (Z3_Interface.Sym (">", _)) [ct, cu] =
       
    95         SOME (mk_lt cu ct)
       
    96     | z3_mk_builtin_fun (Z3_Interface.Sym (">=", _)) [ct, cu] =
       
    97         SOME (mk_le cu ct)
       
    98     | z3_mk_builtin_fun _ _ = NONE
       
    99 in
       
   100 
       
   101 val z3_mk_builtins = {
       
   102   mk_builtin_typ = z3_mk_builtin_typ,
       
   103   mk_builtin_num = z3_mk_builtin_num,
       
   104   mk_builtin_fun = (fn _ => fn sym => fn cts =>
       
   105     (case try (#T o Thm.rep_cterm o hd) cts of
       
   106       SOME @{typ real} => z3_mk_builtin_fun sym cts
       
   107     | _ => NONE)) }
       
   108 
       
   109 end
       
   110 
       
   111 
       
   112 (* Z3 proof reconstruction *)
       
   113 
       
   114 val real_rules = @{lemma
       
   115   "0 + (x::real) = x"
       
   116   "x + 0 = x"
       
   117   "0 * x = 0"
       
   118   "1 * x = x"
       
   119   "x + y = y + x"
       
   120   by auto}
       
   121 
       
   122 val real_linarith_proc = Simplifier.simproc_global @{theory} "fast_real_arith" [
       
   123   "(m::real) < n", "(m::real) <= n", "(m::real) = n"] Lin_Arith.simproc
       
   124 
       
   125 
       
   126 (* setup *)
       
   127 
       
   128 val setup =
       
   129   Context.theory_map (
       
   130     SMTLIB_Interface.add_logic (10, smtlib_logic) #>
       
   131     setup_builtins #>
       
   132     Z3_Interface.add_mk_builtins z3_mk_builtins #>
       
   133     fold Z3_Proof_Reconstruction.add_z3_rule real_rules #>
       
   134     Z3_Proof_Tools.add_simproc real_linarith_proc)
       
   135 
       
   136 end