src/HOL/Tools/SMT/z3_real.ML
author wenzelm
Fri, 12 Nov 2021 13:02:20 +0100
changeset 74766 71a447e4073b
parent 69593 3dda49e08b9d
child 78177 ea7a3cc64df5
permissions -rw-r--r--
clarified properties: avoid empty entry;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57229
diff changeset
     1
(*  Title:      HOL/Tools/SMT/z3_real.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
Z3 setup for reals.
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: 57229
diff changeset
     7
structure Z3_Real: sig end =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     8
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     9
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
    10
fun real_type_parser (SMTLIB.Sym "Real", []) = SOME \<^typ>\<open>Real.real\<close>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    11
  | real_type_parser _ = NONE
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    12
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
    13
fun real_term_parser (SMTLIB.Dec (i, 0), []) = SOME (HOLogic.mk_number \<^typ>\<open>Real.real\<close> i)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57229
diff changeset
    14
  | real_term_parser (SMTLIB.Sym "/", [t1, t2]) =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
    15
      SOME (\<^term>\<open>Rings.divide :: real => _\<close> $ t1 $ t2)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
    16
  | real_term_parser (SMTLIB.Sym "to_real", [t]) = SOME (\<^term>\<open>Int.of_int :: int => _\<close> $ t)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
  | real_term_parser _ = NONE
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    18
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
fun abstract abs t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    20
  (case t of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
    21
    (c as \<^term>\<open>Rings.divide :: real => _\<close>) $ t1 $ t2 =>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    22
      abs t1 ##>> abs t2 #>> (fn (u1, u2) => SOME (c $ u1 $ u2))
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69204
diff changeset
    23
  | (c as \<^term>\<open>Int.of_int :: int => _\<close>) $ t =>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    24
      abs t #>> (fn u => SOME (c $ u))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
  | _ => pair NONE)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    26
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    27
val _ = Theory.setup (Context.theory_map (
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57229
diff changeset
    28
  SMTLIB_Proof.add_type_parser real_type_parser #>
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57229
diff changeset
    29
  SMTLIB_Proof.add_term_parser real_term_parser #>
69204
d5ab1636660b split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 61609
diff changeset
    30
  SMT_Replay_Methods.add_arith_abstracter abstract))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
57229
blanchet
parents: 57222
diff changeset
    32
end;