diff --git a/examples/ml/ml_example.ml b/examples/ml/ml_example.ml
index 31f31d566..823a08fe2 100644
--- a/examples/ml/ml_example.ml
+++ b/examples/ml/ml_example.ml
@@ -26,20 +26,14 @@ exception TestFailedException of string
*)
let model_converter_test ( ctx : context ) =
Printf.printf "ModelConverterTest\n";
- let xr = (arith_expr_of_expr
- (Expr.mk_const ctx (Symbol.mk_string ctx "x")
- (sort_of_arith_sort (arith_sort_of_real_sort (Real.mk_sort ctx))))) in
- let yr = (arith_expr_of_expr
- (Expr.mk_const ctx (Symbol.mk_string ctx "y")
- (sort_of_arith_sort (arith_sort_of_real_sort (Real.mk_sort ctx))))) in
+ let xr = (Expr.mk_const ctx (Symbol.mk_string ctx "x") (Real.mk_sort ctx)) in
+ let yr = (Expr.mk_const ctx (Symbol.mk_string ctx "y") (Real.mk_sort ctx)) in
let g4 = (mk_goal ctx true false false ) in
- (Goal.assert_ g4 [ (mk_gt ctx xr
- (arith_expr_of_real_expr (real_expr_of_rat_num
- (Real.mk_numeral_nd ctx 10 1)))) ]) ;
+ (Goal.assert_ g4 [ (mk_gt ctx xr (Real.mk_numeral_nd ctx 10 1)) ]) ;
(Goal.assert_ g4 [ (mk_eq ctx
- (expr_of_arith_expr yr)
- (expr_of_arith_expr (Arithmetic.mk_add ctx [ xr; (arith_expr_of_real_expr (real_expr_of_rat_num (Real.mk_numeral_nd ctx 1 1))) ]) ) ) ] ) ;
- (Goal.assert_ g4 [ (mk_gt ctx yr (arith_expr_of_real_expr (real_expr_of_rat_num (Real.mk_numeral_nd ctx 1 1)))) ]) ;
+ yr
+ (Arithmetic.mk_add ctx [ xr; (Real.mk_numeral_nd ctx 1 1) ])) ]) ;
+ (Goal.assert_ g4 [ (mk_gt ctx yr (Real.mk_numeral_nd ctx 1 1)) ]) ;
(
let ar = (Tactic.apply (mk_tactic ctx "simplify") g4 None) in
if ((get_num_subgoals ar) == 1 &&
@@ -83,7 +77,7 @@ let basic_tests ( ctx : context ) =
let fname = (mk_string ctx "f") in
let x = (mk_string ctx "x") in
let y = (mk_string ctx "y") in
- let bs = (sort_of_bool_sort (Boolean.mk_sort ctx)) in
+ let bs = (Boolean.mk_sort ctx) in
let domain = [ bs; bs ] in
let f = (FuncDecl.mk_func_decl ctx fname domain bs) in
let fapp = (mk_app ctx f
@@ -123,10 +117,8 @@ let basic_tests ( ctx : context ) =
Printf.printf "Test passed.\n"
);
(Goal.assert_ g [ (mk_eq ctx
- (mk_numeral_int ctx 1
- (sort_of_bitvec_sort (BitVector.mk_sort ctx 32)))
- (mk_numeral_int ctx 2
- (sort_of_bitvec_sort (BitVector.mk_sort ctx 32)))) ] )
+ (mk_numeral_int ctx 1 (BitVector.mk_sort ctx 32))
+ (mk_numeral_int ctx 2 (BitVector.mk_sort ctx 32))) ] )
;
(
let ar = (Tactic.apply (mk_tactic ctx "smt") g None) in
@@ -157,10 +149,10 @@ let basic_tests ( ctx : context ) =
);
(
let g3 = (mk_goal ctx true true false) in
- let xc = (Expr.mk_const ctx (Symbol.mk_string ctx "x") (sort_of_arith_sort (arith_sort_of_int_sort (Integer.mk_sort ctx)))) in
- let yc = (Expr.mk_const ctx (Symbol.mk_string ctx "y") (sort_of_arith_sort (arith_sort_of_int_sort (Integer.mk_sort ctx)))) in
- (Goal.assert_ g3 [ (mk_eq ctx xc (mk_numeral_int ctx 1 (sort_of_arith_sort (arith_sort_of_int_sort (Integer.mk_sort ctx))))) ]) ;
- (Goal.assert_ g3 [ (mk_eq ctx yc (mk_numeral_int ctx 2 (sort_of_arith_sort (arith_sort_of_int_sort (Integer.mk_sort ctx))))) ]) ;
+ let xc = (Expr.mk_const ctx (Symbol.mk_string ctx "x") (Integer.mk_sort ctx)) in
+ let yc = (Expr.mk_const ctx (Symbol.mk_string ctx "y") (Integer.mk_sort ctx)) in
+ (Goal.assert_ g3 [ (mk_eq ctx xc (mk_numeral_int ctx 1 (Integer.mk_sort ctx))) ]) ;
+ (Goal.assert_ g3 [ (mk_eq ctx yc (mk_numeral_int ctx 2 (Integer.mk_sort ctx))) ]) ;
let constr = (mk_eq ctx xc yc) in
(Goal.assert_ g3 [ constr ] ) ;
let ar = (Tactic.apply (mk_tactic ctx "smt") g3 None) in
@@ -228,9 +220,9 @@ let _ =
let rs = (Real.mk_sort ctx) in
Printf.printf "int symbol: %s\n" (Symbol.to_string is);
Printf.printf "string symbol: %s\n" (Symbol.to_string ss);
- Printf.printf "bool sort: %s\n" (Sort.to_string (sort_of_bool_sort bs));
- Printf.printf "int sort: %s\n" (Sort.to_string (sort_of_arith_sort (arith_sort_of_int_sort ints)));
- Printf.printf "real sort: %s\n" (Sort.to_string (sort_of_arith_sort (arith_sort_of_real_sort rs)));
+ Printf.printf "bool sort: %s\n" (Sort.to_string bs);
+ Printf.printf "int sort: %s\n" (Sort.to_string ints);
+ Printf.printf "real sort: %s\n" (Sort.to_string rs);
basic_tests ctx ;
Printf.printf "Disposing...\n";
Gc.full_major ()
diff --git a/src/api/ml/z3.ml b/src/api/ml/z3.ml
index 56d319c1a..0676f1a39 100644
--- a/src/api/ml/z3.ml
+++ b/src/api/ml/z3.ml
@@ -133,52 +133,34 @@ let mk_context ( cfg : ( string * string ) list ) =
module Symbol =
struct
- (* Symbol types *)
- type int_symbol = z3_native_object
- type string_symbol = z3_native_object
-
- type symbol =
- | S_Int of int_symbol
- | S_Str of string_symbol
-
+ type symbol = z3_native_object
let create_i ( ctx : context ) ( no : Z3native.ptr ) =
- let res : int_symbol = { m_ctx = ctx ;
- m_n_obj = null ;
- inc_ref = z3obj_nil_ref ;
- dec_ref = z3obj_nil_ref } in
+ let res : symbol = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = z3obj_nil_ref ;
+ dec_ref = z3obj_nil_ref } in
(z3obj_sno res ctx no) ;
(z3obj_create res) ;
res
-
+
let create_s ( ctx : context ) ( no : Z3native.ptr ) =
- let res : string_symbol = { m_ctx = ctx ;
- m_n_obj = null ;
- inc_ref = z3obj_nil_ref ;
- dec_ref = z3obj_nil_ref } in
+ let res : symbol = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = z3obj_nil_ref ;
+ dec_ref = z3obj_nil_ref } in
(z3obj_sno res ctx no) ;
(z3obj_create res) ;
res
let create ( ctx : context ) ( no : Z3native.ptr ) =
match (symbol_kind_of_int (Z3native.get_symbol_kind (context_gno ctx) no)) with
- | INT_SYMBOL -> S_Int (create_i ctx no)
- | STRING_SYMBOL -> S_Str (create_s ctx no)
+ | INT_SYMBOL -> (create_i ctx no)
+ | STRING_SYMBOL -> (create_s ctx no)
- let gc ( x : symbol ) =
- match x with
- | S_Int(n) -> (z3obj_gc n)
- | S_Str(n) -> (z3obj_gc n)
-
- let gnc ( x : symbol ) =
- match x with
- | S_Int(n) -> (z3obj_gnc n)
- | S_Str(n) -> (z3obj_gnc n)
-
- let gno ( x : symbol ) =
- match x with
- | S_Int(n) -> (z3obj_gno n)
- | S_Str(n) -> (z3obj_gno n)
+ let gc ( x : symbol ) = (z3obj_gc x)
+ let gnc ( x : symbol ) = (z3obj_gnc x)
+ let gno ( x : symbol ) = (z3obj_gno x)
let symbol_lton ( a : symbol list ) =
let f ( e : symbol ) = (gno e) in
@@ -187,18 +169,18 @@ struct
let kind ( o : symbol ) = (symbol_kind_of_int (Z3native.get_symbol_kind (gnc o) (gno o)))
let is_int_symbol ( o : symbol ) = (kind o) == INT_SYMBOL
let is_string_symbol ( o : symbol ) = (kind o) == STRING_SYMBOL
- let get_int (o : int_symbol) = Z3native.get_symbol_int (z3obj_gnc o) (z3obj_gno o)
- let get_string (o : string_symbol) = Z3native.get_symbol_string (z3obj_gnc o) (z3obj_gno o)
+ let get_int (o : symbol) = Z3native.get_symbol_int (z3obj_gnc o) (z3obj_gno o)
+ let get_string (o : symbol ) = Z3native.get_symbol_string (z3obj_gnc o) (z3obj_gno o)
let to_string ( o : symbol ) =
match (kind o) with
| INT_SYMBOL -> (string_of_int (Z3native.get_symbol_int (gnc o) (gno o)))
| STRING_SYMBOL -> (Z3native.get_symbol_string (gnc o) (gno o))
let mk_int ( ctx : context ) ( i : int ) =
- S_Int (create_i ctx (Z3native.mk_int_symbol (context_gno ctx) i))
+ (create_i ctx (Z3native.mk_int_symbol (context_gno ctx) i))
let mk_string ( ctx : context ) ( s : string ) =
- S_Str (create_s ctx (Z3native.mk_string_symbol (context_gno ctx) s))
+ (create_s ctx (Z3native.mk_string_symbol (context_gno ctx) s))
let mk_ints ( ctx : context ) ( names : int list ) =
let f elem = mk_int ( ctx : context ) elem in
@@ -353,7 +335,6 @@ open AST
module Sort =
struct
type sort = Sort of AST.ast
- type uninterpreted_sort = UninterpretedSort of sort
let sort_of_ptr : context -> Z3native.ptr -> sort = fun ctx no ->
let q = (z3_native_object_of_ast_ptr ctx no) in
@@ -373,14 +354,7 @@ struct
| UNKNOWN_SORT -> raise (Z3native.Exception "Unknown sort kind encountered")
let ast_of_sort s = match s with Sort(x) -> x
- let sort_of_uninterpreted_sort s = match s with UninterpretedSort(x) -> x
-
- let uninterpreted_sort_of_sort s = match s with Sort(a) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.UNINTERPRETED_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- UninterpretedSort(s)
-
+
let gc ( x : sort ) = (match x with Sort(a) -> (z3obj_gc a))
let gnc ( x : sort ) = (match x with Sort(a) -> (z3obj_gnc a))
let gno ( x : sort ) = (match x with Sort(a) -> (z3obj_gno a))
@@ -409,7 +383,7 @@ struct
dec_ref = Z3native.dec_ref } in
(z3obj_sno res ctx (Z3native.mk_uninterpreted_sort (context_gno ctx) (Symbol.gno s))) ;
(z3obj_create res) ;
- UninterpretedSort(Sort(res))
+ Sort(res)
let mk_uninterpreted_s ( ctx : context ) ( s : string ) =
mk_uninterpreted ctx (Symbol.mk_string ( ctx : context ) s)
@@ -713,9 +687,9 @@ and Expr :
sig
type expr = Expr of AST.ast
val expr_of_ptr : context -> Z3native.ptr -> expr
- val c_of_expr : expr -> context
- val nc_of_expr : expr -> Z3native.ptr
- val ptr_of_expr : expr -> Z3native.ptr
+ val gc : expr -> context
+ val gnc : expr -> Z3native.ptr
+ val gno : expr -> Z3native.ptr
val expr_lton : expr list -> Z3native.ptr array
val ast_of_expr : expr -> AST.ast
val expr_of_ast : AST.ast -> expr
@@ -762,9 +736,9 @@ sig
end = struct
type expr = Expr of AST.ast
- let c_of_expr e = match e with Expr(a) -> (z3obj_gc a)
- let nc_of_expr e = match e with Expr(a) -> (z3obj_gnc a)
- let ptr_of_expr e = match e with Expr(a) -> (z3obj_gno a)
+ let gc e = match e with Expr(a) -> (z3obj_gc a)
+ let gnc e = match e with Expr(a) -> (z3obj_gnc a)
+ let gno e = match e with Expr(a) -> (z3obj_gno a)
let expr_of_ptr : context -> Z3native.ptr -> expr = fun ctx no ->
if ast_kind_of_int (Z3native.get_ast_kind (context_gno ctx) no) == QUANTIFIER_AST then
@@ -802,60 +776,60 @@ end = struct
expr_of_ptr ctx o
let simplify ( x : expr ) ( p : Params.params option ) = match p with
- | None -> expr_of_ptr (c_of_expr x) (Z3native.simplify (nc_of_expr x) (ptr_of_expr x))
- | Some pp -> expr_of_ptr (c_of_expr x) (Z3native.simplify_ex (nc_of_expr x) (ptr_of_expr x) (z3obj_gno pp))
+ | None -> expr_of_ptr (Expr.gc x) (Z3native.simplify (gnc x) (gno x))
+ | Some pp -> expr_of_ptr (Expr.gc x) (Z3native.simplify_ex (gnc x) (gno x) (z3obj_gno pp))
let get_simplify_help ( ctx : context ) =
Z3native.simplify_get_help (context_gno ctx)
let get_simplify_parameter_descrs ( ctx : context ) =
Params.ParamDescrs.param_descrs_of_ptr ctx (Z3native.simplify_get_param_descrs (context_gno ctx))
- let get_func_decl ( x : expr ) = FuncDecl.func_decl_of_ptr (c_of_expr x) (Z3native.get_app_decl (nc_of_expr x) (ptr_of_expr x))
+ let get_func_decl ( x : expr ) = FuncDecl.func_decl_of_ptr (Expr.gc x) (Z3native.get_app_decl (gnc x) (gno x))
- let get_bool_value ( x : expr ) = lbool_of_int (Z3native.get_bool_value (nc_of_expr x) (ptr_of_expr x))
+ let get_bool_value ( x : expr ) = lbool_of_int (Z3native.get_bool_value (gnc x) (gno x))
- let get_num_args ( x : expr ) = Z3native.get_app_num_args (nc_of_expr x) (ptr_of_expr x)
+ let get_num_args ( x : expr ) = Z3native.get_app_num_args (gnc x) (gno x)
let get_args ( x : expr ) = let n = (get_num_args x) in
- let f i = expr_of_ptr (c_of_expr x) (Z3native.get_app_arg (nc_of_expr x) (ptr_of_expr x) i) in
+ let f i = expr_of_ptr (Expr.gc x) (Z3native.get_app_arg (gnc x) (gno x) i) in
mk_list f n
let update ( x : expr ) ( args : expr list ) =
if (List.length args <> (get_num_args x)) then
raise (Z3native.Exception "Number of arguments does not match")
else
- expr_of_ptr (c_of_expr x) (Z3native.update_term (nc_of_expr x) (ptr_of_expr x) (List.length args) (expr_lton args))
+ expr_of_ptr (Expr.gc x) (Z3native.update_term (gnc x) (gno x) (List.length args) (expr_lton args))
let substitute ( x : expr ) from to_ =
if (List.length from) <> (List.length to_) then
raise (Z3native.Exception "Argument sizes do not match")
else
- expr_of_ptr (c_of_expr x) (Z3native.substitute (nc_of_expr x) (ptr_of_expr x) (List.length from) (expr_lton from) (expr_lton to_))
+ expr_of_ptr (Expr.gc x) (Z3native.substitute (gnc x) (gno x) (List.length from) (expr_lton from) (expr_lton to_))
let substitute_one ( x : expr ) from to_ =
substitute ( x : expr ) [ from ] [ to_ ]
let substitute_vars ( x : expr ) to_ =
- expr_of_ptr (c_of_expr x) (Z3native.substitute_vars (nc_of_expr x) (ptr_of_expr x) (List.length to_) (expr_lton to_))
+ expr_of_ptr (Expr.gc x) (Z3native.substitute_vars (gnc x) (gno x) (List.length to_) (expr_lton to_))
let translate ( x : expr ) to_ctx =
- if (c_of_expr x) == to_ctx then
+ if (Expr.gc x) == to_ctx then
x
else
- expr_of_ptr to_ctx (Z3native.translate (nc_of_expr x) (ptr_of_expr x) (context_gno to_ctx))
+ expr_of_ptr to_ctx (Z3native.translate (gnc x) (gno x) (context_gno to_ctx))
- let to_string ( x : expr ) = Z3native.ast_to_string (nc_of_expr x) (ptr_of_expr x)
+ let to_string ( x : expr ) = Z3native.ast_to_string (gnc x) (gno x)
- let is_numeral ( x : expr ) = (Z3native.is_numeral_ast (nc_of_expr x) (ptr_of_expr x))
+ let is_numeral ( x : expr ) = (Z3native.is_numeral_ast (gnc x) (gno x))
- let is_well_sorted ( x : expr ) = Z3native.is_well_sorted (nc_of_expr x) (ptr_of_expr x)
+ let is_well_sorted ( x : expr ) = Z3native.is_well_sorted (gnc x) (gno x)
- let get_sort ( x : expr ) = sort_of_ptr (c_of_expr x) (Z3native.get_sort (nc_of_expr x) (ptr_of_expr x))
+ let get_sort ( x : expr ) = sort_of_ptr (Expr.gc x) (Z3native.get_sort (gnc x) (gno x))
let is_bool ( x : expr ) = (match x with Expr(a) -> (AST.is_expr a)) &&
- (Z3native.is_eq_sort (nc_of_expr x)
- (Z3native.mk_bool_sort (nc_of_expr x))
- (Z3native.get_sort (nc_of_expr x) (ptr_of_expr x)))
+ (Z3native.is_eq_sort (gnc x)
+ (Z3native.mk_bool_sort (gnc x))
+ (Z3native.get_sort (gnc x) (gno x)))
let is_const ( x : expr ) = (match x with Expr(a) -> (AST.is_expr a)) &&
(get_num_args x) == 0 &&
@@ -901,86 +875,52 @@ open Expr
module Boolean =
struct
- type bool_sort = BoolSort of Sort.sort
- type bool_expr = BoolExpr of Expr.expr
-
- let bool_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- let a = (AST.ast_of_ptr ctx no) in
- BoolExpr(Expr.Expr(a))
-
- let bool_expr_of_expr e =
- match e with Expr.Expr(a) ->
- let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
- if (q != Z3enums.BOOL_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- BoolExpr(e)
-
- let bool_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- BoolSort(sort_of_ptr ctx no)
-
- let sort_of_bool_sort s = match s with BoolSort(x) -> x
-
- let bool_sort_of_sort s = match s with Sort(a) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.BOOL_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- BoolSort(s)
-
- let expr_of_bool_expr e = match e with BoolExpr(x) -> x
-
- let gc ( x : bool_expr ) = match x with BoolExpr(e) -> (Expr.c_of_expr e)
- let gnc ( x : bool_expr ) = match x with BoolExpr(e) -> (Expr.nc_of_expr e)
- let gno ( x : bool_expr ) = match x with BoolExpr(e) -> (Expr.ptr_of_expr e)
-
let mk_sort ( ctx : context ) =
- BoolSort(sort_of_ptr ctx (Z3native.mk_bool_sort (context_gno ctx)))
+ (sort_of_ptr ctx (Z3native.mk_bool_sort (context_gno ctx)))
let mk_const ( ctx : context ) ( name : Symbol.symbol ) =
- let s = (match (mk_sort ctx) with BoolSort(q) -> q) in
- BoolExpr(Expr.mk_const ctx name s)
+ (Expr.mk_const ctx name (mk_sort ctx))
let mk_const_s ( ctx : context ) ( name : string ) =
mk_const ctx (Symbol.mk_string ctx name)
let mk_true ( ctx : context ) =
- bool_expr_of_ptr ctx (Z3native.mk_true (context_gno ctx))
+ expr_of_ptr ctx (Z3native.mk_true (context_gno ctx))
let mk_false ( ctx : context ) =
- bool_expr_of_ptr ctx (Z3native.mk_false (context_gno ctx))
+ expr_of_ptr ctx (Z3native.mk_false (context_gno ctx))
let mk_val ( ctx : context ) ( value : bool ) =
if value then mk_true ctx else mk_false ctx
let mk_eq ( ctx : context ) ( x : expr ) ( y : expr ) =
- bool_expr_of_ptr ctx (Z3native.mk_eq (context_gno ctx) (ptr_of_expr x) (ptr_of_expr y))
+ expr_of_ptr ctx (Z3native.mk_eq (context_gno ctx) (Expr.gno x) (Expr.gno y))
let mk_distinct ( ctx : context ) ( args : expr list ) =
- bool_expr_of_ptr ctx (Z3native.mk_distinct (context_gno ctx) (List.length args) (expr_lton args))
+ expr_of_ptr ctx (Z3native.mk_distinct (context_gno ctx) (List.length args) (expr_lton args))
- let mk_not ( ctx : context ) ( a : bool_expr ) =
- bool_expr_of_ptr ctx (Z3native.mk_not (context_gno ctx) (gno a))
+ let mk_not ( ctx : context ) ( a : expr ) =
+ expr_of_ptr ctx (Z3native.mk_not (context_gno ctx) (gno a))
- let mk_ite ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) ( t3 : bool_expr ) =
- bool_expr_of_ptr ctx (Z3native.mk_ite (context_gno ctx) (gno t1) (gno t2) (gno t3))
+ let mk_ite ( ctx : context ) ( t1 : expr ) ( t2 : expr ) ( t3 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_ite (context_gno ctx) (gno t1) (gno t2) (gno t3))
- let mk_iff ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) =
- bool_expr_of_ptr ctx (Z3native.mk_iff (context_gno ctx) (gno t1) (gno t2))
+ let mk_iff ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_iff (context_gno ctx) (gno t1) (gno t2))
- let mk_implies ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) =
- bool_expr_of_ptr ctx (Z3native.mk_implies (context_gno ctx) (gno t1) (gno t2))
+ let mk_implies ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_implies (context_gno ctx) (gno t1) (gno t2))
- let mk_xor ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) =
- bool_expr_of_ptr ctx (Z3native.mk_xor (context_gno ctx) (gno t1) (gno t2))
+ let mk_xor ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_xor (context_gno ctx) (gno t1) (gno t2))
- let mk_and ( ctx : context ) ( args : bool_expr list ) =
- let f x = (ptr_of_expr (expr_of_bool_expr x)) in
- bool_expr_of_ptr ctx (Z3native.mk_and (context_gno ctx) (List.length args) (Array.of_list (List.map f args)))
+ let mk_and ( ctx : context ) ( args : expr list ) =
+ let f x = (Expr.gno (x)) in
+ expr_of_ptr ctx (Z3native.mk_and (context_gno ctx) (List.length args) (Array.of_list (List.map f args)))
- let mk_or ( ctx : context ) ( args : bool_expr list ) =
- let f x = (ptr_of_expr (expr_of_bool_expr x)) in
- bool_expr_of_ptr ctx (Z3native.mk_or (context_gno ctx) (List.length args) (Array.of_list(List.map f args)))
+ let mk_or ( ctx : context ) ( args : expr list ) =
+ let f x = (Expr.gno (x)) in
+ expr_of_ptr ctx (Z3native.mk_or (context_gno ctx) (List.length args) (Array.of_list(List.map f args)))
end
@@ -998,9 +938,9 @@ struct
else
Quantifier(e)
- let gc ( x : quantifier ) = match (x) with Quantifier(e) -> (c_of_expr e)
- let gnc ( x : quantifier ) = match (x) with Quantifier(e) -> (nc_of_expr e)
- let gno ( x : quantifier ) = match (x) with Quantifier(e) -> (ptr_of_expr e)
+ let gc ( x : quantifier ) = match (x) with Quantifier(e) -> (Expr.gc e)
+ let gnc ( x : quantifier ) = match (x) with Quantifier(e) -> (Expr.gnc e)
+ let gno ( x : quantifier ) = match (x) with Quantifier(e) -> (Expr.gno e)
module Pattern =
struct
@@ -1031,7 +971,7 @@ struct
if not (AST.is_var (match x with Expr.Expr(a) -> a)) then
raise (Z3native.Exception "Term is not a bound variable.")
else
- Z3native.get_index_value (nc_of_expr x) (ptr_of_expr x)
+ Z3native.get_index_value (Expr.gnc x) (Expr.gno x)
let is_universal ( x : quantifier ) =
Z3native.is_quantifier_forall (gnc x) (gno x)
@@ -1067,7 +1007,7 @@ struct
mk_list f n
let get_body ( x : quantifier ) =
- Boolean.bool_expr_of_ptr (gc x) (Z3native.get_quantifier_body (gnc x) (gno x))
+ expr_of_ptr (gc x) (Z3native.get_quantifier_body (gnc x) (gno x))
let mk_bound ( ctx : context ) ( index : int ) ( ty : sort ) =
expr_of_ptr ctx (Z3native.mk_bound (context_gno ctx) index (Sort.gno ty))
@@ -1087,7 +1027,7 @@ struct
(List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
(List.length sorts) (sort_lton sorts)
(Symbol.symbol_lton names)
- (ptr_of_expr body)))
+ (Expr.gno body)))
else
Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_ex (context_gno ctx) true
(match weight with | None -> 1 | Some(x) -> x)
@@ -1097,7 +1037,7 @@ struct
(List.length nopatterns) (expr_lton nopatterns)
(List.length sorts) (sort_lton sorts)
(Symbol.symbol_lton names)
- (ptr_of_expr body)))
+ (Expr.gno body)))
let mk_forall_const ( ctx : context ) ( bound_constants : expr list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
if ((List.length nopatterns) == 0 && quantifier_id == None && skolem_id == None) then
@@ -1105,7 +1045,7 @@ struct
(match weight with | None -> 1 | Some(x) -> x)
(List.length bound_constants) (expr_lton bound_constants)
(List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
- (ptr_of_expr body)))
+ (Expr.gno body)))
else
Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_const_ex (context_gno ctx) true
(match weight with | None -> 1 | Some(x) -> x)
@@ -1114,7 +1054,7 @@ struct
(List.length bound_constants) (expr_lton bound_constants)
(List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
(List.length nopatterns) (expr_lton nopatterns)
- (ptr_of_expr body)))
+ (Expr.gno body)))
let mk_exists ( ctx : context ) ( sorts : sort list ) ( names : Symbol.symbol list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
if (List.length sorts) != (List.length names) then
@@ -1125,7 +1065,7 @@ struct
(List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
(List.length sorts) (sort_lton sorts)
(Symbol.symbol_lton names)
- (ptr_of_expr body)))
+ (Expr.gno body)))
else
Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_ex (context_gno ctx) false
(match weight with | None -> 1 | Some(x) -> x)
@@ -1135,7 +1075,7 @@ struct
(List.length nopatterns) (expr_lton nopatterns)
(List.length sorts) (sort_lton sorts)
(Symbol.symbol_lton names)
- (ptr_of_expr body)))
+ (Expr.gno body)))
let mk_exists_const ( ctx : context ) ( bound_constants : expr list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
if ((List.length nopatterns) == 0 && quantifier_id == None && skolem_id == None) then
@@ -1143,7 +1083,7 @@ struct
(match weight with | None -> 1 | Some(x) -> x)
(List.length bound_constants) (expr_lton bound_constants)
(List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
- (ptr_of_expr body)))
+ (Expr.gno body)))
else
Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_const_ex (context_gno ctx) false
(match weight with | None -> 1 | Some(x) -> x)
@@ -1152,7 +1092,7 @@ struct
(List.length bound_constants) (expr_lton bound_constants)
(List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
(List.length nopatterns) (expr_lton nopatterns)
- (ptr_of_expr body)))
+ (Expr.gno body)))
let mk_quantifier ( ctx : context ) ( universal : bool ) ( sorts : sort list ) ( names : Symbol.symbol list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
if (universal) then
@@ -1170,46 +1110,8 @@ end
module Array_ =
struct
- type array_sort = ArraySort of sort
- type array_expr = ArrayExpr of expr
-
- let array_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- let e = (expr_of_ptr ctx no) in
- ArrayExpr(e)
-
- let array_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- let s = (sort_of_ptr ctx no) in
- ArraySort(s)
-
- let sort_of_array_sort s = match s with ArraySort(x) -> x
-
- let array_sort_of_sort s = match s with Sort(a) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.ARRAY_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- ArraySort(s)
-
- let array_expr_of_expr e =
- match e with Expr(a) ->
- let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
- if (q != Z3enums.ARRAY_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- ArrayExpr(e)
-
- let expr_of_array_expr e = match e with ArrayExpr(x) -> x
-
- let sgc ( x : array_sort ) = match (x) with ArraySort(Sort(s)) -> (z3obj_gc s)
- let sgnc ( x : array_sort ) = match (x) with ArraySort(Sort(s)) -> (z3obj_gnc s)
- let sgno ( x : array_sort ) = match (x) with ArraySort(Sort(s)) -> (z3obj_gno s)
-
- let egc ( x : array_expr ) = match (x) with ArrayExpr(Expr(e)) -> (z3obj_gc e)
- let egnc ( x : array_expr ) = match (x) with ArrayExpr(Expr(e)) -> (z3obj_gnc e)
- let egno ( x : array_expr ) = match (x) with ArrayExpr(Expr(e)) -> (z3obj_gno e)
-
let mk_sort ( ctx : context ) ( domain : sort ) ( range : sort ) =
- array_sort_of_ptr ctx (Z3native.mk_array_sort (context_gno ctx) (Sort.gno domain) (Sort.gno range))
+ sort_of_ptr ctx (Z3native.mk_array_sort (context_gno ctx) (Sort.gno domain) (Sort.gno range))
let is_store ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_STORE)
let is_select ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SELECT)
@@ -1218,48 +1120,40 @@ struct
let is_array_map ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ARRAY_MAP)
let is_as_array ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_AS_ARRAY)
let is_array ( x : expr ) =
- (Z3native.is_app (nc_of_expr x) (ptr_of_expr x)) &&
- ((sort_kind_of_int (Z3native.get_sort_kind (nc_of_expr x) (Z3native.get_sort (nc_of_expr x) (ptr_of_expr x)))) == ARRAY_SORT)
+ (Z3native.is_app (Expr.gnc x) (Expr.gno x)) &&
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == ARRAY_SORT)
- let get_domain ( x : array_sort ) = Sort.sort_of_ptr (sgc x) (Z3native.get_array_sort_domain (sgnc x) (sgno x))
- let get_range ( x : array_sort ) = Sort.sort_of_ptr (sgc x) (Z3native.get_array_sort_range (sgnc x) (sgno x))
+ let get_domain ( x : sort ) = Sort.sort_of_ptr (Sort.gc x) (Z3native.get_array_sort_domain (Sort.gnc x) (Sort.gno x))
+ let get_range ( x : sort ) = Sort.sort_of_ptr (Sort.gc x) (Z3native.get_array_sort_range (Sort.gnc x) (Sort.gno x))
let mk_const ( ctx : context ) ( name : Symbol.symbol ) ( domain : sort ) ( range : sort ) =
- ArrayExpr(Expr.mk_const ctx name (match (mk_sort ctx domain range) with ArraySort(s) -> s))
+ (Expr.mk_const ctx name (mk_sort ctx domain range))
let mk_const_s ( ctx : context ) ( name : string ) ( domain : sort ) ( range : sort ) =
mk_const ctx (Symbol.mk_string ctx name) domain range
- let mk_select ( ctx : context ) ( a : array_expr ) ( i : expr ) =
- array_expr_of_ptr ctx (Z3native.mk_select (context_gno ctx) (egno a) (ptr_of_expr i))
+ let mk_select ( ctx : context ) ( a : expr ) ( i : expr ) =
+ expr_of_ptr ctx (Z3native.mk_select (context_gno ctx) (Expr.gno a) (Expr.gno i))
- let mk_store ( ctx : context ) ( a : array_expr ) ( i : expr ) ( v : expr ) =
- array_expr_of_ptr ctx (Z3native.mk_store (context_gno ctx) (egno a) (ptr_of_expr i) (ptr_of_expr v))
+ let mk_store ( ctx : context ) ( a : expr ) ( i : expr ) ( v : expr ) =
+ expr_of_ptr ctx (Z3native.mk_store (context_gno ctx) (Expr.gno a) (Expr.gno i) (Expr.gno v))
let mk_const_array ( ctx : context ) ( domain : sort ) ( v : expr ) =
- array_expr_of_ptr ctx (Z3native.mk_const_array (context_gno ctx) (Sort.gno domain) (ptr_of_expr v))
+ expr_of_ptr ctx (Z3native.mk_const_array (context_gno ctx) (Sort.gno domain) (Expr.gno v))
- let mk_map ( ctx : context ) ( f : func_decl ) ( args : array_expr list ) =
- let m x = (ptr_of_expr (expr_of_array_expr x)) in
- array_expr_of_ptr ctx (Z3native.mk_map (context_gno ctx) (FuncDecl.gno f) (List.length args) (Array.of_list (List.map m args)))
+ let mk_map ( ctx : context ) ( f : func_decl ) ( args : expr list ) =
+ let m x = (Expr.gno x) in
+ expr_of_ptr ctx (Z3native.mk_map (context_gno ctx) (FuncDecl.gno f) (List.length args) (Array.of_list (List.map m args)))
- let mk_term_array ( ctx : context ) ( arg : array_expr ) =
- array_expr_of_ptr ctx (Z3native.mk_array_default (context_gno ctx) (egno arg))
+ let mk_term_array ( ctx : context ) ( arg : expr ) =
+ expr_of_ptr ctx (Z3native.mk_array_default (context_gno ctx) (Expr.gno arg))
end
module Set =
struct
- type set_sort = SetSort of sort
-
- let set_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- let s = (sort_of_ptr ctx no) in
- SetSort(s)
-
- let sort_of_set_sort s = match s with SetSort(x) -> x
-
let mk_sort ( ctx : context ) ( ty : sort ) =
- set_sort_of_ptr ctx (Z3native.mk_set_sort (context_gno ctx) (Sort.gno ty))
+ sort_of_ptr ctx (Z3native.mk_set_sort (context_gno ctx) (Sort.gno ty))
let is_union ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SET_UNION)
let is_intersect ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SET_INTERSECT)
@@ -1275,10 +1169,10 @@ struct
expr_of_ptr ctx (Z3native.mk_full_set (context_gno ctx) (Sort.gno domain))
let mk_set_add ( ctx : context ) ( set : expr ) ( element : expr ) =
- expr_of_ptr ctx (Z3native.mk_set_add (context_gno ctx) (ptr_of_expr set) (ptr_of_expr element))
+ expr_of_ptr ctx (Z3native.mk_set_add (context_gno ctx) (Expr.gno set) (Expr.gno element))
let mk_del ( ctx : context ) ( set : expr ) ( element : expr ) =
- expr_of_ptr ctx (Z3native.mk_set_del (context_gno ctx) (ptr_of_expr set) (ptr_of_expr element))
+ expr_of_ptr ctx (Z3native.mk_set_del (context_gno ctx) (Expr.gno set) (Expr.gno element))
let mk_union ( ctx : context ) ( args : expr list ) =
expr_of_ptr ctx (Z3native.mk_set_union (context_gno ctx) (List.length args) (expr_lton args))
@@ -1287,53 +1181,37 @@ struct
expr_of_ptr ctx (Z3native.mk_set_intersect (context_gno ctx) (List.length args) (expr_lton args))
let mk_difference ( ctx : context ) ( arg1 : expr ) ( arg2 : expr ) =
- expr_of_ptr ctx (Z3native.mk_set_difference (context_gno ctx) (ptr_of_expr arg1) (ptr_of_expr arg2))
+ expr_of_ptr ctx (Z3native.mk_set_difference (context_gno ctx) (Expr.gno arg1) (Expr.gno arg2))
let mk_complement ( ctx : context ) ( arg : expr ) =
- expr_of_ptr ctx (Z3native.mk_set_complement (context_gno ctx) (ptr_of_expr arg))
+ expr_of_ptr ctx (Z3native.mk_set_complement (context_gno ctx) (Expr.gno arg))
let mk_membership ( ctx : context ) ( elem : expr ) ( set : expr ) =
- expr_of_ptr ctx (Z3native.mk_set_member (context_gno ctx) (ptr_of_expr elem) (ptr_of_expr set))
+ expr_of_ptr ctx (Z3native.mk_set_member (context_gno ctx) (Expr.gno elem) (Expr.gno set))
let mk_subset ( ctx : context ) ( arg1 : expr ) ( arg2 : expr ) =
- expr_of_ptr ctx (Z3native.mk_set_subset (context_gno ctx) (ptr_of_expr arg1) (ptr_of_expr arg2))
+ expr_of_ptr ctx (Z3native.mk_set_subset (context_gno ctx) (Expr.gno arg1) (Expr.gno arg2))
end
module FiniteDomain =
struct
- type finite_domain_sort = FiniteDomainSort of sort
-
- let sort_of_finite_domain_sort s = match s with FiniteDomainSort(x) -> x
-
- let finite_domain_sort_of_sort s = match s with Sort(a) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.FINITE_DOMAIN_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- FiniteDomainSort(s)
-
- let gc ( x : finite_domain_sort ) = match (x) with FiniteDomainSort(Sort(s)) -> (z3obj_gc s)
- let gnc ( x : finite_domain_sort ) = match (x) with FiniteDomainSort(Sort(s)) -> (z3obj_gnc s)
- let gno ( x : finite_domain_sort ) = match (x) with FiniteDomainSort(Sort(s))-> (z3obj_gno s)
-
let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( size : int ) =
- let s = (sort_of_ptr ctx (Z3native.mk_finite_domain_sort (context_gno ctx) (Symbol.gno name) size)) in
- FiniteDomainSort(s)
+ (sort_of_ptr ctx (Z3native.mk_finite_domain_sort (context_gno ctx) (Symbol.gno name) size))
let mk_sort_s ( ctx : context ) ( name : string ) ( size : int ) =
mk_sort ctx (Symbol.mk_string ctx name) size
-
let is_finite_domain ( x : expr ) =
- let nc = (nc_of_expr x) in
- (Z3native.is_app (nc_of_expr x) (ptr_of_expr x)) &&
- (sort_kind_of_int (Z3native.get_sort_kind nc (Z3native.get_sort nc (ptr_of_expr x))) == FINITE_DOMAIN_SORT)
+ let nc = (Expr.gnc x) in
+ (Z3native.is_app (Expr.gnc x) (Expr.gno x)) &&
+ (sort_kind_of_int (Z3native.get_sort_kind nc (Z3native.get_sort nc (Expr.gno x))) == FINITE_DOMAIN_SORT)
let is_lt ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_FD_LT)
- let get_size ( x : finite_domain_sort ) =
- let (r, v) = (Z3native.get_finite_domain_sort_size (gnc x) (gno x)) in
+ let get_size ( x : sort ) =
+ let (r, v) = (Z3native.get_finite_domain_sort_size (Sort.gnc x) (Sort.gno x)) in
if r then v
else raise (Z3native.Exception "Conversion failed.")
end
@@ -1341,29 +1219,10 @@ end
module Relation =
struct
- type relation_sort = RelationSort of sort
-
- let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- let s = (sort_of_ptr ctx no) in
- RelationSort(s)
-
- let sort_of_relation_sort s = match s with RelationSort(x) -> x
-
- let relation_sort_of_sort s = match s with Sort(a) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.RELATION_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- RelationSort(s)
-
- let gc ( x : relation_sort ) = match (x) with RelationSort(Sort(s)) -> (z3obj_gc s)
- let gnc ( x : relation_sort ) = match (x) with RelationSort(Sort(s)) -> (z3obj_gnc s)
- let gno ( x : relation_sort ) = match (x) with RelationSort(Sort(s))-> (z3obj_gno s)
-
-
let is_relation ( x : expr ) =
- let nc = (nc_of_expr x) in
- ((Z3native.is_app (nc_of_expr x) (ptr_of_expr x)) &&
- (sort_kind_of_int (Z3native.get_sort_kind nc (Z3native.get_sort nc (ptr_of_expr x))) == RELATION_SORT))
+ let nc = (Expr.gnc x) in
+ ((Z3native.is_app (Expr.gnc x) (Expr.gno x)) &&
+ (sort_kind_of_int (Z3native.get_sort_kind nc (Z3native.get_sort nc (Expr.gno x))) == RELATION_SORT))
let is_store ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_STORE)
let is_empty ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_EMPTY)
@@ -1379,48 +1238,17 @@ struct
let is_select ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_SELECT)
let is_clone ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_CLONE)
- let get_arity ( x : relation_sort ) = Z3native.get_relation_arity (gnc x) (gno x)
+ let get_arity ( x : sort ) = Z3native.get_relation_arity (Sort.gnc x) (Sort.gno x)
- let get_column_sorts ( x : relation_sort ) =
+ let get_column_sorts ( x : sort ) =
let n = get_arity x in
- let f i = (sort_of_ptr (gc x) (Z3native.get_relation_column (gnc x) (gno x) i)) in
+ let f i = (sort_of_ptr (Sort.gc x) (Z3native.get_relation_column (Sort.gnc x) (Sort.gno x) i)) in
mk_list f n
-
end
module Datatype =
struct
- type datatype_sort = DatatypeSort of sort
- type datatype_expr = DatatypeExpr of expr
-
- let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- let s = (sort_of_ptr ctx no) in
- DatatypeSort(s)
-
- let sort_of_datatype_sort s = match s with DatatypeSort(x) -> x
-
- let datatype_sort_of_sort s = match s with Sort(a) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.DATATYPE_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- DatatypeSort(s)
-
- let datatype_expr_of_expr e =
- match e with Expr(a) ->
- let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
- if (q != Z3enums.DATATYPE_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- DatatypeExpr(e)
-
- let expr_of_datatype_expr e = match e with DatatypeExpr(x) -> x
-
- let sgc ( x : datatype_sort ) = match (x) with DatatypeSort(Sort(s)) -> (z3obj_gc s)
- let sgnc ( x : datatype_sort ) = match (x) with DatatypeSort(Sort(s)) -> (z3obj_gnc s)
- let sgno ( x : datatype_sort ) = match (x) with DatatypeSort(Sort(s))-> (z3obj_gno s)
-
module Constructor =
struct
type constructor = z3_native_object
@@ -1517,24 +1345,24 @@ struct
)
c
- let get_num_constructors ( x : datatype_sort ) = Z3native.get_datatype_sort_num_constructors (sgnc x) (sgno x)
+ let get_num_constructors ( x : sort ) = Z3native.get_datatype_sort_num_constructors (Sort.gnc x) (Sort.gno x)
- let get_constructors ( x : datatype_sort ) =
+ let get_constructors ( x : sort ) =
let n = (get_num_constructors x) in
- let f i = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) i) in
+ let f i = func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor (Sort.gnc x) (Sort.gno x) i) in
mk_list f n
- let get_recognizers ( x : datatype_sort ) =
+ let get_recognizers ( x : sort ) =
let n = (get_num_constructors x) in
- let f i = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) i) in
+ let f i = func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_recognizer (Sort.gnc x) (Sort.gno x) i) in
mk_list f n
- let get_accessors ( x : datatype_sort ) =
+ let get_accessors ( x : sort ) =
let n = (get_num_constructors x) in
let f i = (
- let fd = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) i) in
+ let fd = func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor (Sort.gnc x) (Sort.gno x) i) in
let ds = Z3native.get_domain_size (FuncDecl.gnc fd) (FuncDecl.gno fd) in
- let g j = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor_accessor (sgnc x) (sgno x) i j) in
+ let g j = func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor_accessor (Sort.gnc x) (Sort.gno x) i j) in
mk_list g ds
) in
mk_list f n
@@ -1543,467 +1371,133 @@ end
module Enumeration =
struct
- type enum_sort = EnumSort of sort
-
- let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) ( cdecls : Z3native.z3_func_decl list ) ( tdecls : Z3native.z3_func_decl list ) =
- let s = (sort_of_ptr ctx no) in
- let res = EnumSort(s) in
- res
-
- let sort_of_enum_sort s = match s with EnumSort(x) -> x
-
- let sgc ( x : enum_sort ) = match (x) with EnumSort(Sort(s)) -> (z3obj_gc s)
- let sgnc ( x : enum_sort ) = match (x) with EnumSort(Sort(s)) -> (z3obj_gnc s)
- let sgno ( x : enum_sort ) = match (x) with EnumSort(Sort(s))-> (z3obj_gno s)
-
let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( enum_names : Symbol.symbol list ) =
- let (a, b, c) = (Z3native.mk_enumeration_sort (context_gno ctx) (Symbol.gno name) (List.length enum_names) (Symbol.symbol_lton enum_names)) in
- sort_of_ptr ctx a (list_of_array b) (list_of_array c)
+ let (a, _, _) = (Z3native.mk_enumeration_sort (context_gno ctx) (Symbol.gno name) (List.length enum_names) (Symbol.symbol_lton enum_names)) in
+ sort_of_ptr ctx a
let mk_sort_s ( ctx : context ) ( name : string ) ( enum_names : string list ) =
mk_sort ctx (Symbol.mk_string ctx name) (Symbol.mk_strings ctx enum_names)
- let get_const_decls ( x : enum_sort ) =
- let n = Z3native.get_datatype_sort_num_constructors (sgnc x) (sgno x) in
- let f i = (func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) i)) in
+ let get_const_decls ( x : sort ) =
+ let n = Z3native.get_datatype_sort_num_constructors (Sort.gnc x) (Sort.gno x) in
+ let f i = (func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor (Sort.gnc x) (Sort.gno x) i)) in
mk_list f n
- let get_tester_decls ( x : enum_sort ) =
- let n = Z3native.get_datatype_sort_num_constructors (sgnc x) (sgno x) in
- let f i = (func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) i)) in
+ let get_tester_decls ( x : sort ) =
+ let n = Z3native.get_datatype_sort_num_constructors (Sort.gnc x) (Sort.gno x) in
+ let f i = (func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_recognizer (Sort.gnc x) (Sort.gno x) i)) in
mk_list f n
end
module List_ =
-struct
- type list_sort = ListSort of sort
-
- let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) ( nildecl : Z3native.ptr ) ( is_nildecl : Z3native.ptr ) ( consdecl : Z3native.ptr ) ( is_consdecl : Z3native.ptr ) ( headdecl : Z3native.ptr ) ( taildecl : Z3native.ptr ) =
- let s = (sort_of_ptr ctx no) in
- let res = ListSort(s) in
- res
-
- let sort_of_list_sort s = match s with ListSort(x) -> x
-
- let sgc ( x : list_sort ) = match (x) with ListSort(Sort(s)) -> (z3obj_gc s)
- let sgnc ( x : list_sort ) = match (x) with ListSort(Sort(s)) -> (z3obj_gnc s)
- let sgno ( x : list_sort ) = match (x) with ListSort(Sort(s))-> (z3obj_gno s)
-
+struct
let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( elem_sort : sort ) =
- let (r, a, b, c, d, e, f) = (Z3native.mk_list_sort (context_gno ctx) (Symbol.gno name) (Sort.gno elem_sort)) in
- sort_of_ptr ctx r a b c d e f
+ let (r, _, _, _, _, _, _) = (Z3native.mk_list_sort (context_gno ctx) (Symbol.gno name) (Sort.gno elem_sort)) in
+ sort_of_ptr ctx r
let mk_list_s ( ctx : context ) (name : string) elem_sort =
mk_sort ctx (Symbol.mk_string ctx name) elem_sort
- let get_nil_decl ( x : list_sort ) =
- func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) 0)
+ let get_nil_decl ( x : sort ) =
+ func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor (Sort.gnc x) (Sort.gno x) 0)
- let get_is_nil_decl ( x : list_sort ) =
- func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) 0)
+ let get_is_nil_decl ( x : sort ) =
+ func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_recognizer (Sort.gnc x) (Sort.gno x) 0)
- let get_cons_decl ( x : list_sort ) =
- func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) 1)
+ let get_cons_decl ( x : sort ) =
+ func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor (Sort.gnc x) (Sort.gno x) 1)
- let get_is_cons_decl ( x : list_sort ) =
- func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) 1)
+ let get_is_cons_decl ( x : sort ) =
+ func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_recognizer (Sort.gnc x) (Sort.gno x) 1)
- let get_head_decl ( x : list_sort ) =
- func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor_accessor (sgnc x) (sgno x) 1 0)
+ let get_head_decl ( x : sort ) =
+ func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor_accessor (Sort.gnc x) (Sort.gno x) 1 0)
- let get_tail_decl ( x : list_sort ) =
- func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor_accessor (sgnc x) (sgno x) 1 1)
+ let get_tail_decl ( x : sort ) =
+ func_decl_of_ptr (Sort.gc x) (Z3native.get_datatype_sort_constructor_accessor (Sort.gnc x) (Sort.gno x) 1 1)
- let nil ( x : list_sort ) = expr_of_func_app (sgc x) (get_nil_decl x) []
+ let nil ( x : sort ) = expr_of_func_app (Sort.gc x) (get_nil_decl x) []
end
module Tuple =
struct
- type tuple_sort = TupleSort of sort
-
- let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- let s = (sort_of_ptr ctx no) in
- TupleSort(s)
-
- let sort_of_tuple_sort s = match s with TupleSort(x) -> x
-
- let sgc ( x : tuple_sort ) = match (x) with TupleSort(Sort(s)) -> (z3obj_gc s)
- let sgnc ( x : tuple_sort ) = match (x) with TupleSort(Sort(s)) -> (z3obj_gnc s)
- let sgno ( x : tuple_sort ) = match (x) with TupleSort(Sort(s))-> (z3obj_gno s)
-
let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( field_names : Symbol.symbol list ) ( field_sorts : sort list ) =
let (r, _, _) = (Z3native.mk_tuple_sort (context_gno ctx) (Symbol.gno name) (List.length field_names) (Symbol.symbol_lton field_names) (sort_lton field_sorts)) in
sort_of_ptr ctx r
- let get_mk_decl ( x : tuple_sort ) =
- func_decl_of_ptr (sgc x) (Z3native.get_tuple_sort_mk_decl (sgnc x) (sgno x))
+ let get_mk_decl ( x : sort ) =
+ func_decl_of_ptr (Sort.gc x) (Z3native.get_tuple_sort_mk_decl (Sort.gnc x) (Sort.gno x))
- let get_num_fields ( x : tuple_sort ) = Z3native.get_tuple_sort_num_fields (sgnc x) (sgno x)
+ let get_num_fields ( x : sort ) = Z3native.get_tuple_sort_num_fields (Sort.gnc x) (Sort.gno x)
- let get_field_decls ( x : tuple_sort ) =
+ let get_field_decls ( x : sort ) =
let n = get_num_fields x in
- let f i = func_decl_of_ptr (sgc x) (Z3native.get_tuple_sort_field_decl (sgnc x) (sgno x) i) in
+ let f i = func_decl_of_ptr (Sort.gc x) (Z3native.get_tuple_sort_field_decl (Sort.gnc x) (Sort.gno x) i) in
mk_list f n
end
-module rec Arithmetic :
-sig
- type arith_sort = ArithSort of Sort.sort
- type arith_expr = ArithExpr of Expr.expr
-
- val sort_of_arith_sort : arith_sort -> Sort.sort
- val arith_sort_of_sort : Sort.sort -> arith_sort
- val expr_of_arith_expr : arith_expr -> Expr.expr
- val arith_expr_of_expr : Expr.expr -> arith_expr
+module Arithmetic =
+struct
- module rec Integer :
- sig
- type int_sort = IntSort of arith_sort
- type int_expr = IntExpr of arith_expr
- type int_num = IntNum of int_expr
-
- val int_expr_of_ptr : context -> Z3native.ptr -> int_expr
- val int_num_of_ptr : context -> Z3native.ptr -> int_num
-
- val arith_sort_of_int_sort : Integer.int_sort -> arith_sort
- val int_sort_of_arith_sort : arith_sort -> int_sort
- val arith_expr_of_int_expr : int_expr -> arith_expr
- val int_expr_of_int_num : int_num -> int_expr
- val int_expr_of_arith_expr : arith_expr -> int_expr
- val int_num_of_int_expr : int_expr -> int_num
-
- val mk_sort : context -> int_sort
- val get_int : int_num -> int
- val to_string : int_num -> string
- val mk_int_const : context -> Symbol.symbol -> int_expr
- val mk_int_const_s : context -> string -> int_expr
- val mk_mod : context -> int_expr -> int_expr -> int_expr
- val mk_rem : context -> int_expr -> int_expr -> int_expr
- val mk_int_numeral_s : context -> string -> int_num
- val mk_int_numeral_i : context -> int -> int_num
- val mk_int2real : context -> int_expr -> Real.real_expr
- val mk_int2bv : context -> int -> int_expr -> BitVector.bitvec_expr
- end
- and Real :
- sig
- type real_sort = RealSort of arith_sort
- type real_expr = RealExpr of arith_expr
- type rat_num = RatNum of real_expr
-
- val real_expr_of_ptr : context -> Z3native.ptr -> real_expr
- val rat_num_of_ptr : context -> Z3native.ptr -> rat_num
-
- val arith_sort_of_real_sort : Arithmetic.Real.real_sort -> Arithmetic.arith_sort
- val real_sort_of_arith_sort : Arithmetic.arith_sort -> Arithmetic.Real.real_sort
- val arith_expr_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.arith_expr
- val real_expr_of_rat_num : Arithmetic.Real.rat_num -> Arithmetic.Real.real_expr
- val real_expr_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.Real.real_expr
- val rat_num_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.Real.rat_num
-
- val mk_sort : context -> real_sort
- val get_numerator : rat_num -> Integer.int_num
- val get_denominator : rat_num -> Integer.int_num
- val to_decimal_string : rat_num -> int -> string
- val to_string : rat_num -> string
- val mk_real_const : context -> Symbol.symbol -> real_expr
- val mk_real_const_s : context -> string -> real_expr
- val mk_numeral_nd : context -> int -> int -> rat_num
- val mk_numeral_s : context -> string -> rat_num
- val mk_numeral_i : context -> int -> rat_num
- val mk_is_integer : context -> real_expr -> Boolean.bool_expr
- val mk_real2int : context -> real_expr -> Integer.int_expr
- end
- and AlgebraicNumber :
- sig
- type algebraic_num = AlgebraicNum of arith_expr
-
- val arith_expr_of_algebraic_num : algebraic_num -> arith_expr
- val algebraic_num_of_arith_expr : arith_expr -> algebraic_num
-
- val to_upper : algebraic_num -> int -> Real.rat_num
- val to_lower : algebraic_num -> int -> Real.rat_num
- val to_decimal_string : algebraic_num -> int -> string
- val to_string : algebraic_num -> string
- end
-
- val is_int : Expr.expr -> bool
- val is_arithmetic_numeral : Expr.expr -> bool
- val is_le : Expr.expr -> bool
- val is_ge : Expr.expr -> bool
- val is_lt : Expr.expr -> bool
- val is_gt : Expr.expr -> bool
- val is_add : Expr.expr -> bool
- val is_sub : Expr.expr -> bool
- val is_uminus : Expr.expr -> bool
- val is_mul : Expr.expr -> bool
- val is_div : Expr.expr -> bool
- val is_idiv : Expr.expr -> bool
- val is_remainder : Expr.expr -> bool
- val is_modulus : Expr.expr -> bool
- val is_inttoreal : Expr.expr -> bool
- val is_real_to_int : Expr.expr -> bool
- val is_real_is_int : Expr.expr -> bool
- val is_real : Expr.expr -> bool
- val is_int_numeral : Expr.expr -> bool
- val is_rat_num : Expr.expr -> bool
- val is_algebraic_number : Expr.expr -> bool
- val mk_add : context -> arith_expr list -> arith_expr
- val mk_mul : context -> arith_expr list -> arith_expr
- val mk_sub : context -> arith_expr list -> arith_expr
- val mk_unary_minus : context -> arith_expr -> arith_expr
- val mk_div : context -> arith_expr -> arith_expr -> arith_expr
- val mk_power : context -> arith_expr -> arith_expr -> arith_expr
- val mk_lt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
- val mk_le : context -> arith_expr -> arith_expr -> Boolean.bool_expr
- val mk_gt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
- val mk_ge : context -> arith_expr -> arith_expr -> Boolean.bool_expr
-end = struct
- type arith_sort = ArithSort of sort
- type arith_expr = ArithExpr of expr
-
- let arith_expr_of_expr e =
- match e with Expr(a) ->
- let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
- if (q != Z3enums.INT_SORT && q != Z3enums.REAL_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- ArithExpr(e)
-
- let arith_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- arith_expr_of_expr (expr_of_ptr ctx no)
-
- let sort_of_arith_sort s = match s with ArithSort(x) -> x
- let expr_of_arith_expr e = match e with ArithExpr(x) -> x
-
- let arith_sort_of_sort s = match s with Sort(a) ->
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) in
- if (q != Z3enums.INT_SORT && q != Z3enums.REAL_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- ArithSort(s)
-
- let arith_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- arith_sort_of_sort (sort_of_ptr ctx no)
-
- let sgc ( x : arith_sort ) = match (x) with ArithSort(Sort(s)) -> (z3obj_gc s)
- let sgnc ( x : arith_sort ) = match (x) with ArithSort(Sort(s)) -> (z3obj_gnc s)
- let sgno ( x : arith_sort ) = match (x) with ArithSort(Sort(s)) -> (z3obj_gno s)
- let egc ( x : arith_expr ) = match (x) with ArithExpr(e) -> (c_of_expr e)
- let egnc ( x : arith_expr ) = match (x) with ArithExpr(e) -> (nc_of_expr e)
- let egno ( x : arith_expr ) = match (x) with ArithExpr(e) -> (ptr_of_expr e)
-
- module rec Integer :
- sig
- type int_sort = IntSort of arith_sort
- type int_expr = IntExpr of arith_expr
- type int_num = IntNum of int_expr
-
- val int_expr_of_ptr : context -> Z3native.ptr -> int_expr
- val int_num_of_ptr : context -> Z3native.ptr -> int_num
-
- val arith_sort_of_int_sort : Integer.int_sort -> arith_sort
- val int_sort_of_arith_sort : arith_sort -> int_sort
- val arith_expr_of_int_expr : int_expr -> arith_expr
- val int_expr_of_int_num : int_num -> int_expr
- val int_expr_of_arith_expr : arith_expr -> int_expr
- val int_num_of_int_expr : int_expr -> int_num
-
- val mk_sort : context -> int_sort
- val get_int : int_num -> int
- val to_string : int_num -> string
- val mk_int_const : context -> Symbol.symbol -> int_expr
- val mk_int_const_s : context -> string -> int_expr
- val mk_mod : context -> int_expr -> int_expr -> int_expr
- val mk_rem : context -> int_expr -> int_expr -> int_expr
- val mk_int_numeral_s : context -> string -> int_num
- val mk_int_numeral_i : context -> int -> int_num
- val mk_int2real : context -> int_expr -> Real.real_expr
- val mk_int2bv : context -> int -> int_expr -> BitVector.bitvec_expr
- end = struct
- type int_sort = IntSort of arith_sort
- type int_expr = IntExpr of arith_expr
- type int_num = IntNum of int_expr
-
- let int_expr_of_arith_expr e =
- match e with ArithExpr(Expr(a)) ->
- let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
- if (q != Z3enums.INT_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- IntExpr(e)
-
- let int_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- int_expr_of_arith_expr (arith_expr_of_expr (Expr.expr_of_ptr ctx no))
-
- let int_num_of_int_expr e =
- match e with IntExpr(ArithExpr(Expr(a))) ->
- if (not (Z3native.is_numeral_ast (z3obj_gnc a) (z3obj_gno a))) then
- raise (Z3native.Exception "Invalid coercion")
- else
- IntNum(e)
-
- let int_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- int_num_of_int_expr (int_expr_of_ptr ctx no)
-
- let arith_sort_of_int_sort s = match s with IntSort(x) -> x
- let arith_expr_of_int_expr e = match e with IntExpr(x) -> x
- let int_expr_of_int_num e = match e with IntNum(x) -> x
-
- let int_sort_of_arith_sort s = match s with ArithSort(Sort(a)) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.INT_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- IntSort(s)
-
- let int_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- int_sort_of_arith_sort (arith_sort_of_sort (Sort.sort_of_ptr ctx no))
-
- let sgc ( x : int_sort ) = match (x) with IntSort(s) -> (sgc s)
- let sgnc ( x : int_sort ) = match (x) with IntSort(s) -> (sgnc s)
- let sgno ( x : int_sort ) = match (x) with IntSort(s) -> (sgno s)
- let egc ( x : int_expr ) = match (x) with IntExpr(e) -> (egc e)
- let egnc ( x : int_expr ) = match (x) with IntExpr(e) -> (egnc e)
- let egno ( x : int_expr ) = match (x) with IntExpr(e) -> (egno e)
- let ngc ( x : int_num ) = match (x) with IntNum(e) -> (egc e)
- let ngnc ( x : int_num ) = match (x) with IntNum(e) -> (egnc e)
- let ngno ( x : int_num ) = match (x) with IntNum(e) -> (egno e)
-
+ module Integer =
+ struct
let mk_sort ( ctx : context ) =
- int_sort_of_ptr ctx (Z3native.mk_int_sort (context_gno ctx))
+ sort_of_ptr ctx (Z3native.mk_int_sort (context_gno ctx))
- let get_int ( x : int_num ) =
- let (r, v) = Z3native.get_numeral_int (ngnc x) (ngno x) in
+ let get_int ( x : expr ) =
+ let (r, v) = Z3native.get_numeral_int (Expr.gnc x) (Expr.gno x) in
if r then v
else raise (Z3native.Exception "Conversion failed.")
- let to_string ( x : int_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+ let to_string ( x : expr ) = Z3native.get_numeral_string (Expr.gnc x) (Expr.gno x)
let mk_int_const ( ctx : context ) ( name : Symbol.symbol ) =
- IntExpr(ArithExpr(Expr.mk_const ctx name (match (mk_sort ctx) with IntSort(ArithSort(s)) -> s)))
+ Expr.mk_const ctx name (mk_sort ctx)
let mk_int_const_s ( ctx : context ) ( name : string ) =
mk_int_const ctx (Symbol.mk_string ctx name)
- let mk_mod ( ctx : context ) ( t1 : int_expr ) ( t2 : int_expr ) =
- int_expr_of_ptr ctx (Z3native.mk_mod (context_gno ctx) (egno t1) (egno t2))
+ let mk_mod ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_mod (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
- let mk_rem ( ctx : context ) ( t1 : int_expr ) ( t2 : int_expr ) =
- int_expr_of_ptr ctx (Z3native.mk_rem (context_gno ctx) (egno t1) (egno t2))
+ let mk_rem ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_rem (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
let mk_int_numeral_s ( ctx : context ) ( v : string ) =
- int_num_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (sgno (mk_sort ctx)))
+ expr_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (Sort.gno (mk_sort ctx)))
let mk_int_numeral_i ( ctx : context ) ( v : int ) =
- int_num_of_ptr ctx (Z3native.mk_int (context_gno ctx) v (sgno (mk_sort ctx)))
+ expr_of_ptr ctx (Z3native.mk_int (context_gno ctx) v (Sort.gno (mk_sort ctx)))
- let mk_int2real ( ctx : context ) ( t : int_expr ) =
- Real.real_expr_of_arith_expr (arith_expr_of_expr (Expr.expr_of_ptr ctx (Z3native.mk_int2real (context_gno ctx) (egno t))))
+ let mk_int2real ( ctx : context ) ( t : expr ) =
+ (Expr.expr_of_ptr ctx (Z3native.mk_int2real (context_gno ctx) (Expr.gno t)))
- let mk_int2bv ( ctx : context ) ( n : int ) ( t : int_expr ) =
- BitVector.bitvec_expr_of_expr (Expr.expr_of_ptr ctx (Z3native.mk_int2bv (context_gno ctx) n (egno t)))
+ let mk_int2bv ( ctx : context ) ( n : int ) ( t : expr ) =
+ (Expr.expr_of_ptr ctx (Z3native.mk_int2bv (context_gno ctx) n (Expr.gno t)))
end
- and Real :
- sig
- type real_sort = RealSort of arith_sort
- type real_expr = RealExpr of arith_expr
- type rat_num = RatNum of real_expr
-
- val real_expr_of_ptr : context -> Z3native.ptr -> real_expr
- val rat_num_of_ptr : context -> Z3native.ptr -> rat_num
-
- val arith_sort_of_real_sort : real_sort -> arith_sort
- val real_sort_of_arith_sort : arith_sort -> real_sort
- val arith_expr_of_real_expr : real_expr -> arith_expr
- val real_expr_of_rat_num : rat_num -> real_expr
- val real_expr_of_arith_expr : arith_expr -> real_expr
- val rat_num_of_real_expr : real_expr -> rat_num
-
- val mk_sort : context -> real_sort
- val get_numerator : rat_num -> Integer.int_num
- val get_denominator : rat_num -> Integer.int_num
- val to_decimal_string : rat_num -> int -> string
- val to_string : rat_num -> string
- val mk_real_const : context -> Symbol.symbol -> real_expr
- val mk_real_const_s : context -> string -> real_expr
- val mk_numeral_nd : context -> int -> int -> rat_num
- val mk_numeral_s : context -> string -> rat_num
- val mk_numeral_i : context -> int -> rat_num
- val mk_is_integer : context -> real_expr -> Boolean.bool_expr
- val mk_real2int : context -> real_expr -> Integer.int_expr
- end = struct
- type real_sort = RealSort of arith_sort
- type real_expr = RealExpr of arith_expr
- type rat_num = RatNum of real_expr
-
- let arith_sort_of_real_sort s = match s with RealSort(x) -> x
- let arith_expr_of_real_expr e = match e with RealExpr(x) -> x
- let real_expr_of_rat_num e = match e with RatNum(x) -> x
-
- let real_expr_of_arith_expr e =
- match e with ArithExpr(Expr(a)) ->
- let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
- if (q != Z3enums.REAL_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- RealExpr(e)
-
- let real_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- real_expr_of_arith_expr (arith_expr_of_expr (Expr.expr_of_ptr ctx no))
-
- let rat_num_of_real_expr e =
- match e with RealExpr(ArithExpr(Expr(a))) ->
- if (not (Z3native.is_numeral_ast (z3obj_gnc a) (z3obj_gno a))) then
- raise (Z3native.Exception "Invalid coercion")
- else
- RatNum(e)
-
- let rat_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- rat_num_of_real_expr (real_expr_of_ptr ctx no)
-
- let real_sort_of_arith_sort s = match s with ArithSort(Sort(a)) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.REAL_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- RealSort(s)
-
- let real_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- real_sort_of_arith_sort (arith_sort_of_sort (sort_of_ptr ctx no))
-
- let sgc ( x : real_sort ) = match (x) with RealSort(s) -> (sgc s)
- let sgnc ( x : real_sort ) = match (x) with RealSort(s) -> (sgnc s)
- let sgno ( x : real_sort ) = match (x) with RealSort(s) -> (sgno s)
- let egc ( x : real_expr ) = match (x) with RealExpr(e) -> (egc e)
- let egnc ( x : real_expr ) = match (x) with RealExpr(e) -> (egnc e)
- let egno ( x : real_expr ) = match (x) with RealExpr(e) -> (egno e)
- let ngc ( x : rat_num ) = match (x) with RatNum(e) -> (egc e)
- let ngnc ( x : rat_num ) = match (x) with RatNum(e) -> (egnc e)
- let ngno ( x : rat_num ) = match (x) with RatNum(e) -> (egno e)
-
-
+ module Real =
+ struct
let mk_sort ( ctx : context ) =
- real_sort_of_ptr ctx (Z3native.mk_real_sort (context_gno ctx))
+ sort_of_ptr ctx (Z3native.mk_real_sort (context_gno ctx))
- let get_numerator ( x : rat_num ) =
- Integer.int_num_of_ptr (ngc x) (Z3native.get_numerator (ngnc x) (ngno x))
+ let get_numerator ( x : expr ) =
+ expr_of_ptr (Expr.gc x) (Z3native.get_numerator (Expr.gnc x) (Expr.gno x))
- let get_denominator ( x : rat_num ) =
- Integer.int_num_of_ptr (ngc x) (Z3native.get_denominator (ngnc x) (ngno x))
+ let get_denominator ( x : expr ) =
+ expr_of_ptr (Expr.gc x) (Z3native.get_denominator (Expr.gnc x) (Expr.gno x))
- let to_decimal_string ( x : rat_num ) ( precision : int ) =
- Z3native.get_numeral_decimal_string (ngnc x) (ngno x) precision
+ let to_decimal_string ( x : expr ) ( precision : int ) =
+ Z3native.get_numeral_decimal_string (Expr.gnc x) (Expr.gno x) precision
- let to_string ( x : rat_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+ let to_string ( x : expr ) = Z3native.get_numeral_string (Expr.gnc x) (Expr.gno x)
let mk_real_const ( ctx : context ) ( name : Symbol.symbol ) =
- RealExpr(ArithExpr(Expr.mk_const ctx name (match (mk_sort ctx) with RealSort(ArithSort(s)) -> s)))
+ Expr.mk_const ctx name (mk_sort ctx)
let mk_real_const_s ( ctx : context ) ( name : string ) =
mk_real_const ctx (Symbol.mk_string ctx name)
@@ -2012,67 +1506,38 @@ end = struct
if (den == 0) then
raise (Z3native.Exception "Denominator is zero")
else
- rat_num_of_ptr ctx (Z3native.mk_real (context_gno ctx) num den)
+ expr_of_ptr ctx (Z3native.mk_real (context_gno ctx) num den)
let mk_numeral_s ( ctx : context ) ( v : string ) =
- rat_num_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (sgno (mk_sort ctx)))
+ expr_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (Sort.gno (mk_sort ctx)))
let mk_numeral_i ( ctx : context ) ( v : int ) =
- rat_num_of_ptr ctx (Z3native.mk_int (context_gno ctx) v (sgno (mk_sort ctx)))
+ expr_of_ptr ctx (Z3native.mk_int (context_gno ctx) v (Sort.gno (mk_sort ctx)))
- let mk_is_integer ( ctx : context ) ( t : real_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_is_int (context_gno ctx) (egno t)))
+ let mk_is_integer ( ctx : context ) ( t : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_is_int (context_gno ctx) (Expr.gno t)))
- let mk_real2int ( ctx : context ) ( t : real_expr ) =
- Integer.int_expr_of_arith_expr (arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_real2int (context_gno ctx) (egno t))))
+ let mk_real2int ( ctx : context ) ( t : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_real2int (context_gno ctx) (Expr.gno t)))
end
- and AlgebraicNumber :
- sig
- type algebraic_num = AlgebraicNum of arith_expr
-
- val arith_expr_of_algebraic_num : algebraic_num -> arith_expr
- val algebraic_num_of_arith_expr : arith_expr -> algebraic_num
-
- val to_upper : algebraic_num -> int -> Real.rat_num
- val to_lower : algebraic_num -> int -> Real.rat_num
- val to_decimal_string : algebraic_num -> int -> string
- val to_string : algebraic_num -> string
- end = struct
- type algebraic_num = AlgebraicNum of arith_expr
-
- let arith_expr_of_algebraic_num e = match e with AlgebraicNum(x) -> x
-
- let algebraic_num_of_arith_expr e =
- match e with ArithExpr(Expr(a)) ->
- if (not (Z3native.is_algebraic_number (z3obj_gnc a) (z3obj_gno a))) then
- raise (Z3native.Exception "Invalid coercion")
- else
- AlgebraicNum(e)
-
- let algebraic_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- algebraic_num_of_arith_expr (arith_expr_of_expr (expr_of_ptr ctx no))
-
- let ngc ( x : algebraic_num ) = match (x) with AlgebraicNum(e) -> (egc e)
- let ngnc ( x : algebraic_num ) = match (x) with AlgebraicNum(e) -> (egnc e)
- let ngno ( x : algebraic_num ) = match (x) with AlgebraicNum(e) -> (egno e)
-
-
- let to_upper ( x : algebraic_num ) ( precision : int ) =
- Real.rat_num_of_ptr (ngc x) (Z3native.get_algebraic_number_upper (ngnc x) (ngno x) precision)
+ module AlgebraicNumber =
+ struct
+ let to_upper ( x : expr ) ( precision : int ) =
+ expr_of_ptr (Expr.gc x) (Z3native.get_algebraic_number_upper (Expr.gnc x) (Expr.gno x) precision)
- let to_lower ( x : algebraic_num ) precision =
- Real.rat_num_of_ptr (ngc x) (Z3native.get_algebraic_number_lower (ngnc x) (ngno x) precision)
+ let to_lower ( x : expr ) precision =
+ expr_of_ptr (Expr.gc x) (Z3native.get_algebraic_number_lower (Expr.gnc x) (Expr.gno x) precision)
- let to_decimal_string ( x : algebraic_num ) ( precision : int ) =
- Z3native.get_numeral_decimal_string (ngnc x) (ngno x) precision
+ let to_decimal_string ( x : expr ) ( precision : int ) =
+ Z3native.get_numeral_decimal_string (Expr.gnc x) (Expr.gno x) precision
- let to_string ( x : algebraic_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+ let to_string ( x : expr ) = Z3native.get_numeral_string (Expr.gnc x) (Expr.gno x)
end
let is_int ( x : expr ) =
- (Z3native.is_numeral_ast (nc_of_expr x) (nc_of_expr x)) &&
- ((sort_kind_of_int (Z3native.get_sort_kind (nc_of_expr x) (Z3native.get_sort (nc_of_expr x) (nc_of_expr x)))) == INT_SORT)
+ (Z3native.is_numeral_ast (Expr.gnc x) (Expr.gno x)) &&
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == INT_SORT)
let is_arithmetic_numeral ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ANUM)
@@ -2107,224 +1572,55 @@ end = struct
let is_real_is_int ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_IS_INT)
let is_real ( x : expr ) =
- ((sort_kind_of_int (Z3native.get_sort_kind (nc_of_expr x) (Z3native.get_sort (nc_of_expr x) (nc_of_expr x)))) == REAL_SORT)
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == REAL_SORT)
+
let is_int_numeral ( x : expr ) = (Expr.is_numeral x) && (is_int x)
let is_rat_num ( x : expr ) = (Expr.is_numeral x) && (is_real x)
- let is_algebraic_number ( x : expr ) = Z3native.is_algebraic_number (nc_of_expr x) (nc_of_expr x)
+ let is_algebraic_number ( x : expr ) = Z3native.is_algebraic_number (Expr.gnc x) (Expr.gno x)
- let mk_add ( ctx : context ) ( t : arith_expr list ) =
- let f x = (ptr_of_expr (expr_of_arith_expr x)) in
- arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_add (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
+ let mk_add ( ctx : context ) ( t : expr list ) =
+ let f x = (Expr.gno x) in
+ (expr_of_ptr ctx (Z3native.mk_add (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
- let mk_mul ( ctx : context ) ( t : arith_expr list ) =
- let f x = (ptr_of_expr (expr_of_arith_expr x)) in
- arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_mul (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
+ let mk_mul ( ctx : context ) ( t : expr list ) =
+ let f x = (Expr.gno x) in
+ (expr_of_ptr ctx (Z3native.mk_mul (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
- let mk_sub ( ctx : context ) ( t : arith_expr list ) =
- let f x = (ptr_of_expr (expr_of_arith_expr x)) in
- arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_sub (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
+ let mk_sub ( ctx : context ) ( t : expr list ) =
+ let f x = (Expr.gno x) in
+ (expr_of_ptr ctx (Z3native.mk_sub (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
- let mk_unary_minus ( ctx : context ) ( t : arith_expr ) =
- arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_unary_minus (context_gno ctx) (egno t)))
+ let mk_unary_minus ( ctx : context ) ( t : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_unary_minus (context_gno ctx) (Expr.gno t)))
- let mk_div ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
- arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_div (context_gno ctx) (egno t1) (egno t2)))
+ let mk_div ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_div (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
- let mk_power ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
- arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_power (context_gno ctx) (egno t1) (egno t2)))
+ let mk_power ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_power (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
- let mk_lt ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_lt (context_gno ctx) (egno t1) (egno t2)))
+ let mk_lt ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_lt (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
- let mk_le ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_le (context_gno ctx) (egno t1) (egno t2)))
+ let mk_le ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_le (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
- let mk_gt ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_gt (context_gno ctx) (egno t1) (egno t2)))
+ let mk_gt ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_gt (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
- let mk_ge ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_ge (context_gno ctx) (egno t1) (egno t2)))
+ let mk_ge ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_ge (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
end
-and BitVector :
-sig
- type bitvec_sort = BitVecSort of Sort.sort
- type bitvec_expr = BitVecExpr of Expr.expr
- type bitvec_num = BitVecNum of bitvec_expr
-
- val sort_of_bitvec_sort : BitVector.bitvec_sort -> Sort.sort
- val bitvec_sort_of_sort : Sort.sort -> BitVector.bitvec_sort
- val expr_of_bitvec_expr : BitVector.bitvec_expr -> Expr.expr
- val bitvec_expr_of_bitvec_num : BitVector.bitvec_num -> BitVector.bitvec_expr
- val bitvec_expr_of_expr : Expr.expr -> BitVector.bitvec_expr
- val bitvec_num_of_bitvec_expr : BitVector.bitvec_expr -> BitVector.bitvec_num
-
- val mk_sort : context -> int -> bitvec_sort
- val is_bv : Expr.expr -> bool
- val is_bv_numeral : Expr.expr -> bool
- val is_bv_bit1 : Expr.expr -> bool
- val is_bv_bit0 : Expr.expr -> bool
- val is_bv_uminus : Expr.expr -> bool
- val is_bv_add : Expr.expr -> bool
- val is_bv_sub : Expr.expr -> bool
- val is_bv_mul : Expr.expr -> bool
- val is_bv_sdiv : Expr.expr -> bool
- val is_bv_udiv : Expr.expr -> bool
- val is_bv_SRem : Expr.expr -> bool
- val is_bv_urem : Expr.expr -> bool
- val is_bv_smod : Expr.expr -> bool
- val is_bv_sdiv0 : Expr.expr -> bool
- val is_bv_udiv0 : Expr.expr -> bool
- val is_bv_srem0 : Expr.expr -> bool
- val is_bv_urem0 : Expr.expr -> bool
- val is_bv_smod0 : Expr.expr -> bool
- val is_bv_ule : Expr.expr -> bool
- val is_bv_sle : Expr.expr -> bool
- val is_bv_uge : Expr.expr -> bool
- val is_bv_sge : Expr.expr -> bool
- val is_bv_ult : Expr.expr -> bool
- val is_bv_slt : Expr.expr -> bool
- val is_bv_ugt : Expr.expr -> bool
- val is_bv_sgt : Expr.expr -> bool
- val is_bv_and : Expr.expr -> bool
- val is_bv_or : Expr.expr -> bool
- val is_bv_not : Expr.expr -> bool
- val is_bv_xor : Expr.expr -> bool
- val is_bv_nand : Expr.expr -> bool
- val is_bv_nor : Expr.expr -> bool
- val is_bv_xnor : Expr.expr -> bool
- val is_bv_concat : Expr.expr -> bool
- val is_bv_signextension : Expr.expr -> bool
- val is_bv_zeroextension : Expr.expr -> bool
- val is_bv_extract : Expr.expr -> bool
- val is_bv_repeat : Expr.expr -> bool
- val is_bv_reduceor : Expr.expr -> bool
- val is_bv_reduceand : Expr.expr -> bool
- val is_bv_comp : Expr.expr -> bool
- val is_bv_shiftleft : Expr.expr -> bool
- val is_bv_shiftrightlogical : Expr.expr -> bool
- val is_bv_shiftrightarithmetic : Expr.expr -> bool
- val is_bv_rotateleft : Expr.expr -> bool
- val is_bv_rotateright : Expr.expr -> bool
- val is_bv_rotateleftextended : Expr.expr -> bool
- val is_bv_rotaterightextended : Expr.expr -> bool
- val is_int_to_bv : Expr.expr -> bool
- val is_bv_to_int : Expr.expr -> bool
- val is_bv_carry : Expr.expr -> bool
- val is_bv_xor3 : Expr.expr -> bool
- val get_size : bitvec_sort -> int
- val get_int : bitvec_num -> int
- val to_string : bitvec_num -> string
- val mk_const : context -> Symbol.symbol -> int -> bitvec_expr
- val mk_const_s : context -> string -> int -> bitvec_expr
- val mk_not : context -> bitvec_expr -> Expr.expr
- val mk_redand : context -> bitvec_expr -> Expr.expr
- val mk_redor : context -> bitvec_expr -> Expr.expr
- val mk_and : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_or : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_xor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_nand : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_nor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_xnor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_neg : context -> bitvec_expr -> bitvec_expr
- val mk_add : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_sub : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_mul : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_udiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_sdiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_urem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_srem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_smod : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_ult : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_slt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_ule : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_sle : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_uge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_sge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_ugt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_sgt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_concat : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_extract : context -> int -> int -> bitvec_expr -> bitvec_expr
- val mk_sign_ext : context -> int -> bitvec_expr -> bitvec_expr
- val mk_zero_ext : context -> int -> bitvec_expr -> bitvec_expr
- val mk_repeat : context -> int -> bitvec_expr -> bitvec_expr
- val mk_shl : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_lshr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_ashr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_rotate_left : context -> int -> bitvec_expr -> bitvec_expr
- val mk_rotate_right : context -> int -> bitvec_expr -> bitvec_expr
- val mk_ext_rotate_left : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_ext_rotate_right : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
- val mk_bv2int : context -> bitvec_expr -> bool -> Arithmetic.Integer.int_expr
- val mk_add_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
- val mk_add_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_sub_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_sub_no_underflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
- val mk_sdiv_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_neg_no_overflow : context -> bitvec_expr -> Boolean.bool_expr
- val mk_mul_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
- val mk_mul_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
- val mk_numeral : context -> string -> int -> bitvec_num
-end = struct
- type bitvec_sort = BitVecSort of sort
- type bitvec_expr = BitVecExpr of expr
- type bitvec_num = BitVecNum of bitvec_expr
-
- let sort_of_bitvec_sort s = match s with BitVecSort(x) -> x
-
- let bitvec_sort_of_sort s = match s with Sort(a) ->
- if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.BV_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- BitVecSort(s)
-
- let bitvec_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- bitvec_sort_of_sort (sort_of_ptr ctx no)
-
- let bitvec_expr_of_expr e =
- match e with Expr(a) ->
- let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
- let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
- if (q != Z3enums.BV_SORT) then
- raise (Z3native.Exception "Invalid coercion")
- else
- BitVecExpr(e)
-
- let bitvec_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- bitvec_expr_of_expr (expr_of_ptr ctx no)
-
- let bitvec_num_of_bitvec_expr e =
- match e with BitVecExpr(Expr(a)) ->
- if (not (Z3native.is_numeral_ast (z3obj_gnc a) (z3obj_gno a))) then
- raise (Z3native.Exception "Invalid coercion")
- else
- BitVecNum(e)
-
- let bitvec_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
- bitvec_num_of_bitvec_expr (bitvec_expr_of_expr (expr_of_ptr ctx no))
-
- let expr_of_bitvec_expr e = match e with BitVecExpr(x) -> x
- let bitvec_expr_of_bitvec_num e = match e with BitVecNum(x) -> x
-
-
- let sgc ( x : bitvec_sort ) = match (x) with BitVecSort(s) -> (Sort.gc s)
- let sgnc ( x : bitvec_sort ) = match (x) with BitVecSort(s) -> (Sort.gnc s)
- let sgno ( x : bitvec_sort ) = match (x) with BitVecSort(s) -> (Sort.gno s)
- let egc ( x : bitvec_expr ) = match (x) with BitVecExpr(e) -> (c_of_expr e)
- let egnc ( x : bitvec_expr ) = match (x) with BitVecExpr(e) -> (nc_of_expr e)
- let egno ( x : bitvec_expr ) = match (x) with BitVecExpr(e) -> (ptr_of_expr e)
- let ngc ( x : bitvec_num ) = match (x) with BitVecNum(e) -> (egc e)
- let ngnc ( x : bitvec_num ) = match (x) with BitVecNum(e) -> (egnc e)
- let ngno ( x : bitvec_num ) = match (x) with BitVecNum(e) -> (egno e)
-
-
+module BitVector =
+struct
let mk_sort ( ctx : context ) size =
- bitvec_sort_of_ptr ctx (Z3native.mk_bv_sort (context_gno ctx) size)
+ sort_of_ptr ctx (Z3native.mk_bv_sort (context_gno ctx) size)
let is_bv ( x : expr ) =
- ((sort_kind_of_int (Z3native.get_sort_kind (nc_of_expr x) (Z3native.get_sort (nc_of_expr x) (ptr_of_expr x)))) == BV_SORT)
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == BV_SORT)
let is_bv_numeral ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BNUM)
let is_bv_bit1 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BIT1)
let is_bv_bit0 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BIT0)
@@ -2376,112 +1672,112 @@ end = struct
let is_bv_to_int ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BV2INT)
let is_bv_carry ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_CARRY)
let is_bv_xor3 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_XOR3)
- let get_size (x : bitvec_sort ) = Z3native.get_bv_sort_size (sgnc x) (sgno x)
- let get_int ( x : bitvec_num ) =
- let (r, v) = Z3native.get_numeral_int (ngnc x) (ngno x) in
+ let get_size (x : sort ) = Z3native.get_bv_sort_size (Sort.gnc x) (Sort.gno x)
+ let get_int ( x : expr ) =
+ let (r, v) = Z3native.get_numeral_int (Expr.gnc x) (Expr.gno x) in
if r then v
else raise (Z3native.Exception "Conversion failed.")
- let to_string ( x : bitvec_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+ let to_string ( x : expr ) = Z3native.get_numeral_string (Expr.gnc x) (Expr.gno x)
let mk_const ( ctx : context ) ( name : Symbol.symbol ) ( size : int ) =
- BitVecExpr(Expr.mk_const ctx name (match (BitVector.mk_sort ctx size) with BitVecSort(s) -> s))
+ Expr.mk_const ctx name (mk_sort ctx size)
let mk_const_s ( ctx : context ) ( name : string ) ( size : int ) =
mk_const ctx (Symbol.mk_string ctx name) size
- let mk_not ( ctx : context ) ( t : bitvec_expr ) =
- expr_of_ptr ctx (Z3native.mk_bvnot (context_gno ctx) (egno t))
- let mk_redand ( ctx : context ) ( t : bitvec_expr) =
- expr_of_ptr ctx (Z3native.mk_bvredand (context_gno ctx) (egno t))
- let mk_redor ( ctx : context ) ( t : bitvec_expr) =
- expr_of_ptr ctx (Z3native.mk_bvredor (context_gno ctx) (egno t))
- let mk_and ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvand (context_gno ctx) (egno t1) (egno t2))
- let mk_or ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvor (context_gno ctx) (egno t1) (egno t2))
- let mk_xor ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvxor (context_gno ctx) (egno t1) (egno t2))
- let mk_nand ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvnand (context_gno ctx) (egno t1) (egno t2))
- let mk_nor ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvnor (context_gno ctx) (egno t1) (egno t2))
- let mk_xnor ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvxnor (context_gno ctx) (egno t1) (egno t2))
- let mk_neg ( ctx : context ) ( t : bitvec_expr) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvneg (context_gno ctx) (egno t))
- let mk_add ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvadd (context_gno ctx) (egno t1) (egno t2))
- let mk_sub ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvsub (context_gno ctx) (egno t1) (egno t2))
- let mk_mul ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvmul (context_gno ctx) (egno t1) (egno t2))
- let mk_udiv ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvudiv (context_gno ctx) (egno t1) (egno t2))
- let mk_sdiv ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvsdiv (context_gno ctx) (egno t1) (egno t2))
- let mk_urem ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvurem (context_gno ctx) (egno t1) (egno t2))
- let mk_srem ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvsrem (context_gno ctx) (egno t1) (egno t2))
- let mk_smod ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvsmod (context_gno ctx) (egno t1) (egno t2))
- let mk_ult ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvult (context_gno ctx) (egno t1) (egno t2)))
- let mk_slt ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvslt (context_gno ctx) (egno t1) (egno t2)))
- let mk_ule ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvule (context_gno ctx) (egno t1) (egno t2)))
- let mk_sle ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsle (context_gno ctx) (egno t1) (egno t2)))
- let mk_uge ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvuge (context_gno ctx) (egno t1) (egno t2)))
- let mk_sge ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsge (context_gno ctx) (egno t1) (egno t2)))
- let mk_ugt ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvugt (context_gno ctx) (egno t1) (egno t2)))
- let mk_sgt ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsgt (context_gno ctx) (egno t1) (egno t2)))
- let mk_concat ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_concat (context_gno ctx) (egno t1) (egno t2))
- let mk_extract ( ctx : context ) ( high : int ) ( low : int ) ( t : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_extract (context_gno ctx) high low (egno t))
- let mk_sign_ext ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_sign_ext (context_gno ctx) i (egno t))
- let mk_zero_ext ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_zero_ext (context_gno ctx) i (egno t))
- let mk_repeat ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_repeat (context_gno ctx) i (egno t))
- let mk_shl ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvshl (context_gno ctx) (egno t1) (egno t2))
- let mk_lshr ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvlshr (context_gno ctx) (egno t1) (egno t2))
- let mk_ashr ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_bvashr (context_gno ctx) (egno t1) (egno t2))
- let mk_rotate_left ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_rotate_left (context_gno ctx) i (egno t))
- let mk_rotate_right ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_rotate_right (context_gno ctx) i (egno t))
- let mk_ext_rotate_left ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_ext_rotate_left (context_gno ctx) (egno t1) (egno t2))
- let mk_ext_rotate_right ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- bitvec_expr_of_ptr ctx (Z3native.mk_ext_rotate_right (context_gno ctx) (egno t1) (egno t2))
- let mk_bv2int ( ctx : context ) ( t : bitvec_expr ) ( signed : bool ) =
- Arithmetic.Integer.int_expr_of_ptr ctx (Z3native.mk_bv2int (context_gno ctx) (egno t) signed)
- let mk_add_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) ( signed : bool) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvadd_no_overflow (context_gno ctx) (egno t1) (egno t2) signed))
- let mk_add_no_underflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvadd_no_underflow (context_gno ctx) (egno t1) (egno t2)))
- let mk_sub_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsub_no_overflow (context_gno ctx) (egno t1) (egno t2)))
- let mk_sub_no_underflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) ( signed : bool) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsub_no_underflow (context_gno ctx) (egno t1) (egno t2) signed))
- let mk_sdiv_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsdiv_no_overflow (context_gno ctx) (egno t1) (egno t2)))
- let mk_neg_no_overflow ( ctx : context ) ( t : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvneg_no_overflow (context_gno ctx) (egno t)))
- let mk_mul_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) ( signed : bool) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvmul_no_overflow (context_gno ctx) (egno t1) (egno t2) signed))
- let mk_mul_no_underflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvmul_no_underflow (context_gno ctx) (egno t1) (egno t2)))
+ let mk_not ( ctx : context ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvnot (context_gno ctx) (Expr.gno t))
+ let mk_redand ( ctx : context ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvredand (context_gno ctx) (Expr.gno t))
+ let mk_redor ( ctx : context ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvredor (context_gno ctx) (Expr.gno t))
+ let mk_and ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvand (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_or ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvor (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_xor ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvxor (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_nand ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvnand (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_nor ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvnor (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_xnor ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvxnor (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_neg ( ctx : context ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvneg (context_gno ctx) (Expr.gno t))
+ let mk_add ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvadd (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_sub ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvsub (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_mul ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvmul (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_udiv ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvudiv (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_sdiv ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvsdiv (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_urem ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvurem (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_srem ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvsrem (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_smod ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvsmod (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_ult ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvult (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_slt ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvslt (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_ule ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvule (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_sle ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvsle (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_uge ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvuge (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_sge ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvsge (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_ugt ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvugt (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_sgt ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvsgt (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_concat ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_concat (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_extract ( ctx : context ) ( high : int ) ( low : int ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_extract (context_gno ctx) high low (Expr.gno t))
+ let mk_sign_ext ( ctx : context ) ( i : int ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_sign_ext (context_gno ctx) i (Expr.gno t))
+ let mk_zero_ext ( ctx : context ) ( i : int ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_zero_ext (context_gno ctx) i (Expr.gno t))
+ let mk_repeat ( ctx : context ) ( i : int ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_repeat (context_gno ctx) i (Expr.gno t))
+ let mk_shl ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvshl (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_lshr ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvlshr (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_ashr ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvashr (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_rotate_left ( ctx : context ) ( i : int ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_rotate_left (context_gno ctx) i (Expr.gno t))
+ let mk_rotate_right ( ctx : context ) ( i : int ) ( t : expr ) =
+ expr_of_ptr ctx (Z3native.mk_rotate_right (context_gno ctx) i (Expr.gno t))
+ let mk_ext_rotate_left ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_ext_rotate_left (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_ext_rotate_right ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_ext_rotate_right (context_gno ctx) (Expr.gno t1) (Expr.gno t2))
+ let mk_bv2int ( ctx : context ) ( t : expr ) ( signed : bool ) =
+ expr_of_ptr ctx (Z3native.mk_bv2int (context_gno ctx) (Expr.gno t) signed)
+ let mk_add_no_overflow ( ctx : context ) ( t1 : expr ) ( t2 : expr ) ( signed : bool) =
+ (expr_of_ptr ctx (Z3native.mk_bvadd_no_overflow (context_gno ctx) (Expr.gno t1) (Expr.gno t2) signed))
+ let mk_add_no_underflow ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvadd_no_underflow (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_sub_no_overflow ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvsub_no_overflow (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_sub_no_underflow ( ctx : context ) ( t1 : expr ) ( t2 : expr ) ( signed : bool) =
+ (expr_of_ptr ctx (Z3native.mk_bvsub_no_underflow (context_gno ctx) (Expr.gno t1) (Expr.gno t2) signed))
+ let mk_sdiv_no_overflow ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvsdiv_no_overflow (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
+ let mk_neg_no_overflow ( ctx : context ) ( t : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvneg_no_overflow (context_gno ctx) (Expr.gno t)))
+ let mk_mul_no_overflow ( ctx : context ) ( t1 : expr ) ( t2 : expr ) ( signed : bool) =
+ (expr_of_ptr ctx (Z3native.mk_bvmul_no_overflow (context_gno ctx) (Expr.gno t1) (Expr.gno t2) signed))
+ let mk_mul_no_underflow ( ctx : context ) ( t1 : expr ) ( t2 : expr ) =
+ (expr_of_ptr ctx (Z3native.mk_bvmul_no_underflow (context_gno ctx) (Expr.gno t1) (Expr.gno t2)))
let mk_numeral ( ctx : context ) ( v : string ) ( size : int) =
- bitvec_num_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (sgno (BitVector.mk_sort ctx size)))
+ expr_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (Sort.gno (mk_sort ctx size)))
end
@@ -2556,8 +1852,8 @@ struct
let is_garbage ( x : goal ) =
(get_precision x) == GOAL_UNDER_OVER
- let assert_ ( x : goal ) ( constraints : Boolean.bool_expr list ) =
- let f e = Z3native.goal_assert (z3obj_gnc x) (z3obj_gno x) (Boolean.gno e) in
+ let assert_ ( x : goal ) ( constraints : expr list ) =
+ let f e = Z3native.goal_assert (z3obj_gnc x) (z3obj_gno x) (Expr.gno e) in
ignore (List.map f constraints) ;
()
@@ -2572,7 +1868,7 @@ struct
let get_formulas ( x : goal ) =
let n = get_size x in
- let f i = (Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x)
+ let f i = ((expr_of_ptr (z3obj_gc x)
(Z3native.goal_formula (z3obj_gnc x) (z3obj_gno x) i))) in
mk_list f n
@@ -2749,7 +2045,7 @@ struct
exception ModelEvaluationFailedException of string
let eval ( x : model ) ( t : expr ) ( completion : bool ) =
- let (r, v) = (Z3native.model_eval (z3obj_gnc x) (z3obj_gno x) (ptr_of_expr t) completion) in
+ let (r, v) = (Z3native.model_eval (z3obj_gnc x) (z3obj_gno x) (Expr.gno t) completion) in
if not r then
raise (ModelEvaluationFailedException "evaluation failed")
else
@@ -3076,19 +2372,19 @@ struct
let reset ( x : solver ) = Z3native.solver_reset (z3obj_gnc x) (z3obj_gno x)
- let assert_ ( x : solver ) ( constraints : Boolean.bool_expr list ) =
- let f e = (Z3native.solver_assert (z3obj_gnc x) (z3obj_gno x) (Boolean.gno e)) in
+ let assert_ ( x : solver ) ( constraints : expr list ) =
+ let f e = (Z3native.solver_assert (z3obj_gnc x) (z3obj_gno x) (Expr.gno e)) in
ignore (List.map f constraints)
- let assert_and_track_a ( x : solver ) ( cs : Boolean.bool_expr list ) ( ps : Boolean.bool_expr list ) =
+ let assert_and_track_a ( x : solver ) ( cs : expr list ) ( ps : expr list ) =
if ((List.length cs) != (List.length ps)) then
raise (Z3native.Exception "Argument size mismatch")
else
- let f a b = (Z3native.solver_assert_and_track (z3obj_gnc x) (z3obj_gno x) (Boolean.gno a) (Boolean.gno b)) in
+ let f a b = (Z3native.solver_assert_and_track (z3obj_gnc x) (z3obj_gno x) (Expr.gno a) (Expr.gno b)) in
ignore (List.iter2 f cs ps)
- let assert_and_track ( x : solver ) ( c : Boolean.bool_expr ) ( p : Boolean.bool_expr ) =
- Z3native.solver_assert_and_track (z3obj_gnc x) (z3obj_gno x) (Boolean.gno c) (Boolean.gno p)
+ let assert_and_track ( x : solver ) ( c : expr ) ( p : expr ) =
+ Z3native.solver_assert_and_track (z3obj_gnc x) (z3obj_gno x) (Expr.gno c) (Expr.gno p)
let get_num_assertions ( x : solver ) =
let a = AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.solver_get_assertions (z3obj_gnc x) (z3obj_gno x)) in
@@ -3097,15 +2393,15 @@ struct
let get_assertions ( x : solver ) =
let a = AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.solver_get_assertions (z3obj_gnc x) (z3obj_gno x)) in
let n = (AST.ASTVector.get_size a) in
- let f i = Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get a i))) in
+ let f i = (expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get a i))) in
mk_list f n
- let check ( x : solver ) ( assumptions : Boolean.bool_expr list ) =
+ let check ( x : solver ) ( assumptions : expr list ) =
let r =
if ((List.length assumptions) == 0) then
lbool_of_int (Z3native.solver_check (z3obj_gnc x) (z3obj_gno x))
else
- let f x = (ptr_of_expr (Boolean.expr_of_bool_expr x)) in
+ let f x = (Expr.gno x) in
lbool_of_int (Z3native.solver_check_assumptions (z3obj_gnc x) (z3obj_gno x) (List.length assumptions) (Array.of_list (List.map f assumptions)))
in
match r with
@@ -3179,24 +2475,24 @@ struct
let get_param_descrs ( x : fixedpoint ) =
Params.ParamDescrs.param_descrs_of_ptr (z3obj_gc x) (Z3native.fixedpoint_get_param_descrs (z3obj_gnc x) (z3obj_gno x))
- let assert_ ( x : fixedpoint ) ( constraints : Boolean.bool_expr list ) =
- let f e = (Z3native.fixedpoint_assert (z3obj_gnc x) (z3obj_gno x) (Boolean.gno e)) in
+ let assert_ ( x : fixedpoint ) ( constraints : expr list ) =
+ let f e = (Z3native.fixedpoint_assert (z3obj_gnc x) (z3obj_gno x) (Expr.gno e)) in
ignore (List.map f constraints) ;
()
let register_relation ( x : fixedpoint ) ( f : func_decl ) =
Z3native.fixedpoint_register_relation (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno f)
- let add_rule ( x : fixedpoint ) ( rule : Boolean.bool_expr ) ( name : Symbol.symbol option ) =
+ let add_rule ( x : fixedpoint ) ( rule : expr ) ( name : Symbol.symbol option ) =
match name with
- | None -> Z3native.fixedpoint_add_rule (z3obj_gnc x) (z3obj_gno x) (Boolean.gno rule) null
- | Some(y) -> Z3native.fixedpoint_add_rule (z3obj_gnc x) (z3obj_gno x) (Boolean.gno rule) (Symbol.gno y)
+ | None -> Z3native.fixedpoint_add_rule (z3obj_gnc x) (z3obj_gno x) (Expr.gno rule) null
+ | Some(y) -> Z3native.fixedpoint_add_rule (z3obj_gnc x) (z3obj_gno x) (Expr.gno rule) (Symbol.gno y)
let add_fact ( x : fixedpoint ) ( pred : func_decl ) ( args : int list ) =
Z3native.fixedpoint_add_fact (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno pred) (List.length args) (Array.of_list args)
- let query ( x : fixedpoint ) ( query : Boolean.bool_expr ) =
- match (lbool_of_int (Z3native.fixedpoint_query (z3obj_gnc x) (z3obj_gno x) (Boolean.gno query))) with
+ let query ( x : fixedpoint ) ( query : expr ) =
+ match (lbool_of_int (Z3native.fixedpoint_query (z3obj_gnc x) (z3obj_gno x) (Expr.gno query))) with
| L_TRUE -> Solver.SATISFIABLE
| L_FALSE -> Solver.UNSATISFIABLE
| _ -> Solver.UNKNOWN
@@ -3214,8 +2510,8 @@ struct
let pop ( x : fixedpoint ) =
Z3native.fixedpoint_pop (z3obj_gnc x) (z3obj_gno x)
- let update_rule ( x : fixedpoint ) ( rule : Boolean.bool_expr ) ( name : Symbol.symbol ) =
- Z3native.fixedpoint_update_rule (z3obj_gnc x) (z3obj_gno x) (Boolean.gno rule) (Symbol.gno name)
+ let update_rule ( x : fixedpoint ) ( rule : expr ) ( name : Symbol.symbol ) =
+ Z3native.fixedpoint_update_rule (z3obj_gnc x) (z3obj_gno x) (Expr.gno rule) (Symbol.gno name)
let get_answer ( x : fixedpoint ) =
let q = (Z3native.fixedpoint_get_answer (z3obj_gnc x) (z3obj_gno x)) in
@@ -3238,27 +2534,27 @@ struct
Some (expr_of_ptr (z3obj_gc x) q)
let add_cover ( x : fixedpoint ) ( level : int ) ( predicate : func_decl ) ( property : expr ) =
- Z3native.fixedpoint_add_cover (z3obj_gnc x) (z3obj_gno x) level (FuncDecl.gno predicate) (ptr_of_expr property)
+ Z3native.fixedpoint_add_cover (z3obj_gnc x) (z3obj_gno x) level (FuncDecl.gno predicate) (Expr.gno property)
let to_string ( x : fixedpoint ) = Z3native.fixedpoint_to_string (z3obj_gnc x) (z3obj_gno x) 0 [||]
let set_predicate_representation ( x : fixedpoint ) ( f : func_decl ) ( kinds : Symbol.symbol list ) =
Z3native.fixedpoint_set_predicate_representation (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno f) (List.length kinds) (Symbol.symbol_lton kinds)
- let to_string_q ( x : fixedpoint ) ( queries : Boolean.bool_expr list ) =
- let f x = ptr_of_expr (Boolean.expr_of_bool_expr x) in
+ let to_string_q ( x : fixedpoint ) ( queries : expr list ) =
+ let f x = Expr.gno x in
Z3native.fixedpoint_to_string (z3obj_gnc x) (z3obj_gno x) (List.length queries) (Array.of_list (List.map f queries))
let get_rules ( x : fixedpoint ) =
let v = (AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.fixedpoint_get_rules (z3obj_gnc x) (z3obj_gno x))) in
let n = (AST.ASTVector.get_size v) in
- let f i = Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get v i))) in
+ let f i =(expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get v i))) in
mk_list f n
let get_assertions ( x : fixedpoint ) =
let v = (AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.fixedpoint_get_assertions (z3obj_gnc x) (z3obj_gno x))) in
let n = (AST.ASTVector.get_size v) in
- let f i = Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get v i))) in
+ let f i =(expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get v i))) in
mk_list f n
let mk_fixedpoint ( ctx : context ) = create ctx
@@ -3287,10 +2583,10 @@ end
module SMT =
struct
- let benchmark_to_smtstring ( ctx : context ) ( name : string ) ( logic : string ) ( status : string ) ( attributes : string ) ( assumptions : Boolean.bool_expr list ) ( formula : Boolean.bool_expr ) =
+ let benchmark_to_smtstring ( ctx : context ) ( name : string ) ( logic : string ) ( status : string ) ( attributes : string ) ( assumptions : expr list ) ( formula : expr ) =
Z3native.benchmark_to_smtlib_string (context_gno ctx) name logic status attributes
- (List.length assumptions) (let f x = ptr_of_expr (Boolean.expr_of_bool_expr x) in (Array.of_list (List.map f assumptions)))
- (Boolean.gno formula)
+ (List.length assumptions) (let f x = Expr.gno (x) in (Array.of_list (List.map f assumptions)))
+ (Expr.gno formula)
let parse_smtlib_string ( ctx : context ) ( str : string ) ( sort_names : Symbol.symbol list ) ( sorts : sort list ) ( decl_names : Symbol.symbol list ) ( decls : func_decl list ) =
let csn = (List.length sort_names) in
@@ -3328,14 +2624,14 @@ struct
let get_smtlib_formulas ( ctx : context ) =
let n = (get_num_smtlib_formulas ctx ) in
- let f i = Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.get_smtlib_formula (context_gno ctx) i)) in
+ let f i =(expr_of_ptr ctx (Z3native.get_smtlib_formula (context_gno ctx) i)) in
mk_list f n
let get_num_smtlib_assumptions ( ctx : context ) = Z3native.get_smtlib_num_assumptions (context_gno ctx)
let get_smtlib_assumptions ( ctx : context ) =
let n = (get_num_smtlib_assumptions ctx ) in
- let f i = Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.get_smtlib_assumption (context_gno ctx) i)) in
+ let f i = (expr_of_ptr ctx (Z3native.get_smtlib_assumption (context_gno ctx) i)) in
mk_list f n
let get_num_smtlib_decls ( ctx : context ) = Z3native.get_smtlib_num_decls (context_gno ctx)
@@ -3360,14 +2656,14 @@ struct
if (csn != cs || cdn != cd) then
raise (Z3native.Exception "Argument size mismatch")
else
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.parse_smtlib2_string (context_gno ctx) str
- cs
- (Symbol.symbol_lton sort_names)
- (sort_lton sorts)
- cd
- (Symbol.symbol_lton decl_names)
- (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))))
-
+ (expr_of_ptr ctx (Z3native.parse_smtlib2_string (context_gno ctx) str
+ cs
+ (Symbol.symbol_lton sort_names)
+ (sort_lton sorts)
+ cd
+ (Symbol.symbol_lton decl_names)
+ (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))))
+
let parse_smtlib2_file ( ctx : context ) ( file_name : string ) ( sort_names : Symbol.symbol list ) ( sorts : sort list ) ( decl_names : Symbol.symbol list ) ( decls : func_decl list ) =
let csn = (List.length sort_names) in
let cs = (List.length sorts) in
@@ -3376,13 +2672,13 @@ struct
if (csn != cs || cdn != cd) then
raise (Z3native.Exception "Argument size mismatch")
else
- Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.parse_smtlib2_string (context_gno ctx) file_name
- cs
- (Symbol.symbol_lton sort_names)
- (sort_lton sorts)
- cd
- (Symbol.symbol_lton decl_names)
- (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))))
+ (expr_of_ptr ctx (Z3native.parse_smtlib2_string (context_gno ctx) file_name
+ cs
+ (Symbol.symbol_lton sort_names)
+ (sort_lton sorts)
+ cd
+ (Symbol.symbol_lton decl_names)
+ (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))))
end
diff --git a/src/api/ml/z3.mli b/src/api/ml/z3.mli
index 9f451c45f..908b3023d 100644
--- a/src/api/ml/z3.mli
+++ b/src/api/ml/z3.mli
@@ -70,14 +70,7 @@ end
(** Symbols are used to name several term and type constructors *)
module Symbol :
sig
- (** Numbered Symbols *)
- type int_symbol
-
- (** Named Symbols *)
- type string_symbol
-
- (** Symbols *)
- type symbol = S_Int of int_symbol | S_Str of string_symbol
+ type symbol
(** The kind of the symbol (int or string) *)
val kind : symbol -> Z3enums.symbol_kind
@@ -89,10 +82,10 @@ sig
val is_string_symbol : symbol -> bool
(** The int value of the symbol. *)
- val get_int : int_symbol -> int
+ val get_int : symbol -> int
(** The string value of the symbol. *)
- val get_string : string_symbol -> string
+ val get_string : symbol -> string
(** A string representation of the symbol. *)
val to_string : symbol -> string
@@ -251,15 +244,9 @@ end
(** The Sort module implements type information for ASTs *)
module Sort :
sig
- (** Sorts *)
type sort = Sort of AST.ast
- (** Uninterpreted Sorts *)
- type uninterpreted_sort = UninterpretedSort of sort
-
val ast_of_sort : sort -> AST.ast
- val sort_of_uninterpreted_sort : uninterpreted_sort -> sort
- val uninterpreted_sort_of_sort : sort -> uninterpreted_sort
(** Comparison operator.
@return True if the two sorts are from the same context
@@ -279,10 +266,10 @@ sig
val to_string : sort -> string
(** Create a new uninterpreted sort. *)
- val mk_uninterpreted : context -> Symbol.symbol -> uninterpreted_sort
+ val mk_uninterpreted : context -> Symbol.symbol -> sort
(** Create a new uninterpreted sort. *)
- val mk_uninterpreted_s : context -> string -> uninterpreted_sort
+ val mk_uninterpreted_s : context -> string -> sort
end
(** Function declarations *)
@@ -597,58 +584,50 @@ end
(** Boolean expressions *)
module Boolean :
sig
- type bool_sort = BoolSort of Sort.sort
- type bool_expr = BoolExpr of Expr.expr
-
- val expr_of_bool_expr : bool_expr -> Expr.expr
- val sort_of_bool_sort : bool_sort -> Sort.sort
- val bool_sort_of_sort : Sort.sort -> bool_sort
- val bool_expr_of_expr : Expr.expr -> bool_expr
-
(** Create a Boolean sort *)
- val mk_sort : context -> bool_sort
+ val mk_sort : context -> Sort.sort
(** Create a Boolean constant. *)
- val mk_const : context -> Symbol.symbol -> bool_expr
+ val mk_const : context -> Symbol.symbol -> Expr.expr
(** Create a Boolean constant. *)
- val mk_const_s : context -> string -> bool_expr
+ val mk_const_s : context -> string -> Expr.expr
(** The true Term. *)
- val mk_true : context -> bool_expr
+ val mk_true : context -> Expr.expr
(** The false Term. *)
- val mk_false : context -> bool_expr
+ val mk_false : context -> Expr.expr
(** Creates a Boolean value. *)
- val mk_val : context -> bool -> bool_expr
+ val mk_val : context -> bool -> Expr.expr
(** Creates the equality between two expr's. *)
- val mk_eq : context -> Expr.expr -> Expr.expr -> bool_expr
+ val mk_eq : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Creates a distinct term. *)
- val mk_distinct : context -> Expr.expr list -> bool_expr
+ val mk_distinct : context -> Expr.expr list -> Expr.expr
(** Mk an expression representing not(a). *)
- val mk_not : context -> bool_expr -> bool_expr
+ val mk_not : context -> Expr.expr -> Expr.expr
(** Create an expression representing an if-then-else: ite(t1, t2, t3). *)
- val mk_ite : context -> bool_expr -> bool_expr -> bool_expr -> bool_expr
+ val mk_ite : context -> Expr.expr -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 iff t2. *)
- val mk_iff : context -> bool_expr -> bool_expr -> bool_expr
+ val mk_iff : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 -> t2. *)
- val mk_implies : context -> bool_expr -> bool_expr -> bool_expr
+ val mk_implies : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 xor t2. *)
- val mk_xor : context -> bool_expr -> bool_expr -> bool_expr
+ val mk_xor : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing the AND of args *)
- val mk_and : context -> bool_expr list -> bool_expr
+ val mk_and : context -> Expr.expr list -> Expr.expr
(** Create an expression representing the OR of args *)
- val mk_or : context -> bool_expr list -> bool_expr
+ val mk_or : context -> Expr.expr list -> Expr.expr
end
(** Quantifier expressions *)
@@ -731,7 +710,7 @@ sig
val get_bound_variable_sorts : quantifier -> Sort.sort list
(** The body of the quantifier. *)
- val get_body : quantifier -> Boolean.bool_expr
+ val get_body : quantifier -> Expr.expr
(** Creates a new bound variable. *)
val mk_bound : context -> int -> Sort.sort -> Expr.expr
@@ -761,17 +740,8 @@ end
(** Functions to manipulate Array expressions *)
module Array_ :
sig
- type array_sort = ArraySort of Sort.sort
- type array_expr = ArrayExpr of Expr.expr
-
- val sort_of_array_sort : array_sort -> Sort.sort
- val array_sort_of_sort : Sort.sort -> array_sort
- val expr_of_array_expr : array_expr -> Expr.expr
-
- val array_expr_of_expr : Expr.expr -> array_expr
-
(** Create a new array sort. *)
- val mk_sort : context -> Sort.sort -> Sort.sort -> array_sort
+ val mk_sort : context -> Sort.sort -> Sort.sort -> Sort.sort
(** Indicates whether the term is an array store.
It satisfies select(store(a,i,v),j) = if i = j then v else select(a,j).
@@ -802,16 +772,16 @@ sig
val is_array : Expr.expr -> bool
(** The domain of the array sort. *)
- val get_domain : array_sort -> Sort.sort
+ val get_domain : Sort.sort -> Sort.sort
(** The range of the array sort. *)
- val get_range : array_sort -> Sort.sort
+ val get_range : Sort.sort -> Sort.sort
(** Create an array constant. *)
- val mk_const : context -> Symbol.symbol -> Sort.sort -> Sort.sort -> array_expr
+ val mk_const : context -> Symbol.symbol -> Sort.sort -> Sort.sort -> Expr.expr
(** Create an array constant. *)
- val mk_const_s : context -> string -> Sort.sort -> Sort.sort -> array_expr
+ val mk_const_s : context -> string -> Sort.sort -> Sort.sort -> Expr.expr
(** Array read.
@@ -823,7 +793,7 @@ sig
The sort of the result is range.
{!Array_.mk_sort}
{!mk_store} *)
- val mk_select : context -> array_expr -> Expr.expr -> array_expr
+ val mk_select : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Array update.
@@ -839,7 +809,7 @@ sig
respect to i may be a different value).
{!Array_.mk_sort}
{!mk_select} *)
- val mk_store : context -> array_expr -> Expr.expr -> Expr.expr -> array_expr
+ val mk_store : context -> Expr.expr -> Expr.expr -> Expr.expr -> Expr.expr
(** Create a constant array.
@@ -847,7 +817,7 @@ sig
produces the value v.
{!Array_.mk_sort}
{!mk_select} *)
- val mk_const_array : context -> Sort.sort -> Expr.expr -> array_expr
+ val mk_const_array : context -> Sort.sort -> Expr.expr -> Expr.expr
(** Maps f on the argument arrays.
@@ -857,24 +827,20 @@ sig
{!Array_.mk_sort}
{!mk_select}
{!mk_store} *)
- val mk_map : context -> FuncDecl.func_decl -> array_expr list -> array_expr
+ val mk_map : context -> FuncDecl.func_decl -> Expr.expr list -> Expr.expr
(** Access the array default value.
Produces the default range value, for arrays that can be represented as
finite maps with a default range value. *)
- val mk_term_array : context -> array_expr -> array_expr
+ val mk_term_array : context -> Expr.expr -> Expr.expr
end
(** Functions to manipulate Set expressions *)
module Set :
sig
- type set_sort = SetSort of Sort.sort
-
- val sort_of_set_sort : set_sort -> Sort.sort
-
(** Create a set type. *)
- val mk_sort : context -> Sort.sort -> set_sort
+ val mk_sort : context -> Sort.sort -> Sort.sort
(** Indicates whether the term is set union *)
val is_union : Expr.expr -> bool
@@ -925,16 +891,11 @@ end
(** Functions to manipulate Finite Domain expressions *)
module FiniteDomain :
sig
- type finite_domain_sort = FiniteDomainSort of Sort.sort
-
- val sort_of_finite_domain_sort : finite_domain_sort -> Sort.sort
- val finite_domain_sort_of_sort : Sort.sort -> finite_domain_sort
+ (** Create a new finite domain sort. *)
+ val mk_sort : context -> Symbol.symbol -> int -> Sort.sort
(** Create a new finite domain sort. *)
- val mk_sort : context -> Symbol.symbol -> int -> finite_domain_sort
-
- (** Create a new finite domain sort. *)
- val mk_sort_s : context -> string -> int -> finite_domain_sort
+ val mk_sort_s : context -> string -> int -> Sort.sort
(** Indicates whether the term is of an array sort. *)
val is_finite_domain : Expr.expr -> bool
@@ -943,18 +904,13 @@ sig
val is_lt : Expr.expr -> bool
(** The size of the finite domain sort. *)
- val get_size : finite_domain_sort -> int
+ val get_size : Sort.sort -> int
end
(** Functions to manipulate Relation expressions *)
module Relation :
sig
- type relation_sort = RelationSort of Sort.sort
-
- val sort_of_relation_sort : relation_sort -> Sort.sort
- val relation_sort_of_sort : Sort.sort -> relation_sort
-
(** Indicates whether the term is of a relation sort. *)
val is_relation : Expr.expr -> bool
@@ -1034,23 +990,15 @@ sig
val is_clone : Expr.expr -> bool
(** The arity of the relation sort. *)
- val get_arity : relation_sort -> int
+ val get_arity : Sort.sort -> int
(** The sorts of the columns of the relation sort. *)
- val get_column_sorts : relation_sort -> relation_sort list
+ val get_column_sorts : Sort.sort -> Sort.sort list
end
(** Functions to manipulate Datatype expressions *)
module Datatype :
sig
- type datatype_sort = DatatypeSort of Sort.sort
- type datatype_expr = DatatypeExpr of Expr.expr
-
- val sort_of_datatype_sort : datatype_sort -> Sort.sort
- val datatype_sort_of_sort : Sort.sort -> datatype_sort
- val expr_of_datatype_expr : datatype_expr -> Expr.expr
- val datatype_expr_of_expr : Expr.expr -> datatype_expr
-
(** Datatype Constructors *)
module Constructor :
sig
@@ -1080,160 +1028,129 @@ sig
val mk_constructor_s : context -> string -> Symbol.symbol -> Symbol.symbol list -> Sort.sort list -> int list -> Constructor.constructor
(** Create a new datatype sort. *)
- val mk_sort : context -> Symbol.symbol -> Constructor.constructor list -> datatype_sort
+ val mk_sort : context -> Symbol.symbol -> Constructor.constructor list -> Sort.sort
(** Create a new datatype sort. *)
- val mk_sort_s : context -> string -> Constructor.constructor list -> datatype_sort
+ val mk_sort_s : context -> string -> Constructor.constructor list -> Sort.sort
(** Create mutually recursive datatypes. *)
- val mk_sorts : context -> Symbol.symbol list -> Constructor.constructor list list -> datatype_sort list
+ val mk_sorts : context -> Symbol.symbol list -> Constructor.constructor list list -> Sort.sort list
(** Create mutually recursive data-types. *)
- val mk_sorts_s : context -> string list -> Constructor.constructor list list -> datatype_sort list
+ val mk_sorts_s : context -> string list -> Constructor.constructor list list -> Sort.sort list
(** The number of constructors of the datatype sort. *)
- val get_num_constructors : datatype_sort -> int
+ val get_num_constructors : Sort.sort -> int
(** The constructors. *)
- val get_constructors : datatype_sort -> FuncDecl.func_decl list
+ val get_constructors : Sort.sort -> FuncDecl.func_decl list
(** The recognizers. *)
- val get_recognizers : datatype_sort -> FuncDecl.func_decl list
+ val get_recognizers : Sort.sort -> FuncDecl.func_decl list
(** The constructor accessors. *)
- val get_accessors : datatype_sort -> FuncDecl.func_decl list list
+ val get_accessors : Sort.sort -> FuncDecl.func_decl list list
end
(** Functions to manipulate Enumeration expressions *)
module Enumeration :
sig
- type enum_sort = EnumSort of Sort.sort
-
- val sort_of_enum_sort : enum_sort -> Sort.sort
+ (** Create a new enumeration sort. *)
+ val mk_sort : context -> Symbol.symbol -> Symbol.symbol list -> Sort.sort
(** Create a new enumeration sort. *)
- val mk_sort : context -> Symbol.symbol -> Symbol.symbol list -> enum_sort
-
- (** Create a new enumeration sort. *)
- val mk_sort_s : context -> string -> string list -> enum_sort
+ val mk_sort_s : context -> string -> string list -> Sort.sort
(** The function declarations of the constants in the enumeration. *)
- val get_const_decls : enum_sort -> FuncDecl.func_decl list
+ val get_const_decls : Sort.sort -> FuncDecl.func_decl list
(** The test predicates for the constants in the enumeration. *)
- val get_tester_decls : enum_sort -> FuncDecl.func_decl list
+ val get_tester_decls : Sort.sort -> FuncDecl.func_decl list
end
(** Functions to manipulate List expressions *)
module List_ :
sig
- type list_sort = ListSort of Sort.sort
-
- val sort_of_list_sort : list_sort -> Sort.sort
+ (** Create a new list sort. *)
+ val mk_sort : context -> Symbol.symbol -> Sort.sort -> Sort.sort
(** Create a new list sort. *)
- val mk_sort : context -> Symbol.symbol -> Sort.sort -> list_sort
-
- (** Create a new list sort. *)
- val mk_list_s : context -> string -> Sort.sort -> list_sort
+ val mk_list_s : context -> string -> Sort.sort -> Sort.sort
(** The declaration of the nil function of this list sort. *)
- val get_nil_decl : list_sort -> FuncDecl.func_decl
+ val get_nil_decl : Sort.sort -> FuncDecl.func_decl
(** The declaration of the isNil function of this list sort. *)
- val get_is_nil_decl : list_sort -> FuncDecl.func_decl
+ val get_is_nil_decl : Sort.sort -> FuncDecl.func_decl
(** The declaration of the cons function of this list sort. *)
- val get_cons_decl : list_sort -> FuncDecl.func_decl
+ val get_cons_decl : Sort.sort -> FuncDecl.func_decl
(** The declaration of the isCons function of this list sort. *)
- val get_is_cons_decl : list_sort -> FuncDecl.func_decl
+ val get_is_cons_decl : Sort.sort -> FuncDecl.func_decl
(** The declaration of the head function of this list sort. *)
- val get_head_decl : list_sort -> FuncDecl.func_decl
+ val get_head_decl : Sort.sort -> FuncDecl.func_decl
(** The declaration of the tail function of this list sort. *)
- val get_tail_decl : list_sort -> FuncDecl.func_decl
+ val get_tail_decl : Sort.sort -> FuncDecl.func_decl
(** The empty list. *)
- val nil : list_sort -> Expr.expr
+ val nil : Sort.sort -> Expr.expr
end
(** Functions to manipulate Tuple expressions *)
module Tuple :
sig
- type tuple_sort = TupleSort of Sort.sort
-
- val sort_of_tuple_sort : tuple_sort -> Sort.sort
-
(** Create a new tuple sort. *)
- val mk_sort : context -> Symbol.symbol -> Symbol.symbol list -> Sort.sort list -> tuple_sort
+ val mk_sort : context -> Symbol.symbol -> Symbol.symbol list -> Sort.sort list -> Sort.sort
(** The constructor function of the tuple. *)
- val get_mk_decl : tuple_sort -> FuncDecl.func_decl
+ val get_mk_decl : Sort.sort -> FuncDecl.func_decl
(** The number of fields in the tuple. *)
- val get_num_fields : tuple_sort -> int
+ val get_num_fields : Sort.sort -> int
(** The field declarations. *)
- val get_field_decls : tuple_sort -> FuncDecl.func_decl list
+ val get_field_decls : Sort.sort -> FuncDecl.func_decl list
end
(** Functions to manipulate arithmetic expressions *)
module rec Arithmetic :
sig
- type arith_sort = ArithSort of Sort.sort
- type arith_expr = ArithExpr of Expr.expr
-
- val sort_of_arith_sort : Arithmetic.arith_sort -> Sort.sort
- val arith_sort_of_sort : Sort.sort -> Arithmetic.arith_sort
- val expr_of_arith_expr : Arithmetic.arith_expr -> Expr.expr
- val arith_expr_of_expr : Expr.expr -> Arithmetic.arith_expr
-
(** Integer Arithmetic *)
module rec Integer :
sig
- type int_sort = IntSort of arith_sort
- type int_expr = IntExpr of arith_expr
- type int_num = IntNum of int_expr
-
- val arith_sort_of_int_sort : Arithmetic.Integer.int_sort -> Arithmetic.arith_sort
- val int_sort_of_arith_sort : Arithmetic.arith_sort -> Arithmetic.Integer.int_sort
- val arith_expr_of_int_expr : Arithmetic.Integer.int_expr -> Arithmetic.arith_expr
- val int_expr_of_int_num : Arithmetic.Integer.int_num -> Arithmetic.Integer.int_expr
- val int_expr_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.Integer.int_expr
- val int_num_of_int_expr : Arithmetic.Integer.int_expr -> Arithmetic.Integer.int_num
-
(** Create a new integer sort. *)
- val mk_sort : context -> int_sort
+ val mk_sort : context -> Sort.sort
(** Retrieve the int value. *)
- val get_int : int_num -> int
+ val get_int : Expr.expr -> int
(** Returns a string representation of the numeral. *)
- val to_string : int_num -> string
+ val to_string : Expr.expr -> string
(** Creates an integer constant. *)
- val mk_int_const : context -> Symbol.symbol -> int_expr
+ val mk_int_const : context -> Symbol.symbol -> Expr.expr
(** Creates an integer constant. *)
- val mk_int_const_s : context -> string -> int_expr
+ val mk_int_const_s : context -> string -> Expr.expr
(** Create an expression representing t1 mod t2.
The arguments must have int type. *)
- val mk_mod : context -> int_expr -> int_expr -> int_expr
+ val mk_mod : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 rem t2.
The arguments must have int type. *)
- val mk_rem : context -> int_expr -> int_expr -> int_expr
+ val mk_rem : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an integer numeral. *)
- val mk_int_numeral_s : context -> string -> int_num
+ val mk_int_numeral_s : context -> string -> Expr.expr
(** Create an integer numeral.
@return A Term with the given value and sort Integer *)
- val mk_int_numeral_i : context -> int -> int_num
+ val mk_int_numeral_i : context -> int -> Expr.expr
(** Coerce an integer to a real.
@@ -1243,7 +1160,7 @@ sig
You can take the floor of a real by creating an auxiliary integer Term k and
and asserting MakeInt2Real(k) <= t1 < MkInt2Real(k)+1.
The argument must be of integer sort. *)
- val mk_int2real : context -> int_expr -> Real.real_expr
+ val mk_int2real : context -> Expr.expr -> Expr.expr
(** Create an n-bit bit-vector from an integer argument.
@@ -1253,94 +1170,78 @@ sig
when solving constraints with this function.
The argument must be of integer sort. *)
- val mk_int2bv : context -> int -> int_expr -> BitVector.bitvec_expr
+ val mk_int2bv : context -> int -> Expr.expr -> Expr.expr
end
(** Real Arithmetic *)
and Real :
sig
- type real_sort = RealSort of arith_sort
- type real_expr = RealExpr of arith_expr
- type rat_num = RatNum of real_expr
-
- val arith_sort_of_real_sort : Arithmetic.Real.real_sort -> Arithmetic.arith_sort
- val real_sort_of_arith_sort : Arithmetic.arith_sort -> Arithmetic.Real.real_sort
- val arith_expr_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.arith_expr
- val real_expr_of_rat_num : Arithmetic.Real.rat_num -> Arithmetic.Real.real_expr
- val real_expr_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.Real.real_expr
- val rat_num_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.Real.rat_num
-
(** Create a real sort. *)
- val mk_sort : context -> real_sort
+ val mk_sort : context -> Sort.sort
(** The numerator of a rational numeral. *)
- val get_numerator : rat_num -> Integer.int_num
+ val get_numerator : Expr.expr-> Expr.expr
(** The denominator of a rational numeral. *)
- val get_denominator : rat_num -> Integer.int_num
+ val get_denominator : Expr.expr-> Expr.expr
(** Returns a string representation in decimal notation.
The result has at most as many decimal places as indicated by the int argument.*)
- val to_decimal_string : rat_num -> int -> string
+ val to_decimal_string : Expr.expr-> int -> string
(** Returns a string representation of the numeral. *)
- val to_string : rat_num -> string
+ val to_string : Expr.expr-> string
(** Creates a real constant. *)
- val mk_real_const : context -> Symbol.symbol -> real_expr
+ val mk_real_const : context -> Symbol.symbol -> Expr.expr
(** Creates a real constant. *)
- val mk_real_const_s : context -> string -> real_expr
+ val mk_real_const_s : context -> string -> Expr.expr
(** Create a real numeral from a fraction.
@return A Term with rational value and sort Real
{!mk_numeral_s} *)
- val mk_numeral_nd : context -> int -> int -> rat_num
+ val mk_numeral_nd : context -> int -> int -> Expr.expr
(** Create a real numeral.
@return A Term with the given value and sort Real *)
- val mk_numeral_s : context -> string -> rat_num
+ val mk_numeral_s : context -> string -> Expr.expr
(** Create a real numeral.
@return A Term with the given value and sort Real *)
- val mk_numeral_i : context -> int -> rat_num
+ val mk_numeral_i : context -> int -> Expr.expr
(** Creates an expression that checks whether a real number is an integer. *)
- val mk_is_integer : context -> real_expr -> Boolean.bool_expr
+ val mk_is_integer : context -> Expr.expr -> Expr.expr
(** Coerce a real to an integer.
The semantics of this function follows the SMT-LIB standard for the function to_int.
The argument must be of real sort. *)
- val mk_real2int : context -> real_expr -> Integer.int_expr
+ val mk_real2int : context -> Expr.expr -> Expr.expr
end
(** Algebraic Numbers *)
and AlgebraicNumber :
sig
- type algebraic_num = AlgebraicNum of arith_expr
-
- val arith_expr_of_algebraic_num : Arithmetic.AlgebraicNumber.algebraic_num -> Arithmetic.arith_expr
- val algebraic_num_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.AlgebraicNumber.algebraic_num
-
(** Return a upper bound for a given real algebraic number.
The interval isolating the number is smaller than 1/10^precision.
{!is_algebraic_number}
@return A numeral Expr of sort Real *)
- val to_upper : algebraic_num -> int -> Real.rat_num
+ val to_upper : Expr.expr -> int -> Expr.expr
(** Return a lower bound for the given real algebraic number.
The interval isolating the number is smaller than 1/10^precision.
{!is_algebraic_number}
@return A numeral Expr of sort Real *)
- val to_lower : algebraic_num -> int -> Real.rat_num
+ val to_lower : Expr.expr -> int -> Expr.expr
(** Returns a string representation in decimal notation.
The result has at most as many decimal places as the int argument provided.*)
- val to_decimal_string : algebraic_num -> int -> string
+ val to_decimal_string : Expr.expr -> int -> string
(** Returns a string representation of the numeral. *)
- val to_string : algebraic_num -> string
+ val to_string : Expr.expr -> string
end
(** Indicates whether the term is of integer sort. *)
@@ -1407,52 +1308,41 @@ sig
val is_algebraic_number : Expr.expr -> bool
(** Create an expression representing t[0] + t[1] + .... *)
- val mk_add : context -> arith_expr list -> arith_expr
+ val mk_add : context -> Expr.expr list -> Expr.expr
(** Create an expression representing t[0] * t[1] * .... *)
- val mk_mul : context -> arith_expr list -> arith_expr
+ val mk_mul : context -> Expr.expr list -> Expr.expr
(** Create an expression representing t[0] - t[1] - .... *)
- val mk_sub : context -> arith_expr list -> arith_expr
+ val mk_sub : context -> Expr.expr list -> Expr.expr
(** Create an expression representing -t. *)
- val mk_unary_minus : context -> arith_expr -> arith_expr
+ val mk_unary_minus : context -> Expr.expr -> Expr.expr
(** Create an expression representing t1 / t2. *)
- val mk_div : context -> arith_expr -> arith_expr -> arith_expr
+ val mk_div : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 ^ t2. *)
- val mk_power : context -> arith_expr -> arith_expr -> arith_expr
+ val mk_power : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 < t2 *)
- val mk_lt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+ val mk_lt : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 <= t2 *)
- val mk_le : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+ val mk_le : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 > t2 *)
- val mk_gt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+ val mk_gt : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an expression representing t1 >= t2 *)
- val mk_ge : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+ val mk_ge : context -> Expr.expr -> Expr.expr -> Expr.expr
end
(** Functions to manipulate bit-vector expressions *)
and BitVector :
sig
- type bitvec_sort = BitVecSort of Sort.sort
- type bitvec_expr = BitVecExpr of Expr.expr
- type bitvec_num = BitVecNum of bitvec_expr
-
- val sort_of_bitvec_sort : BitVector.bitvec_sort -> Sort.sort
- val bitvec_sort_of_sort : Sort.sort -> BitVector.bitvec_sort
- val expr_of_bitvec_expr : BitVector.bitvec_expr -> Expr.expr
- val bitvec_expr_of_bitvec_num : BitVector.bitvec_num -> BitVector.bitvec_expr
- val bitvec_expr_of_expr : Expr.expr -> BitVector.bitvec_expr
- val bitvec_num_of_bitvec_expr : BitVector.bitvec_expr -> BitVector.bitvec_num
-
(** Create a new bit-vector sort. *)
- val mk_sort : context -> int -> bitvec_sort
+ val mk_sort : context -> int -> Sort.sort
(** Indicates whether the terms is of bit-vector sort. *)
val is_bv : Expr.expr -> bool
@@ -1624,71 +1514,71 @@ sig
val is_bv_xor3 : Expr.expr -> bool
(** The size of a bit-vector sort. *)
- val get_size : bitvec_sort -> int
+ val get_size : Sort.sort -> int
(** Retrieve the int value. *)
- val get_int : bitvec_num -> int
+ val get_int : Expr.expr -> int
(** Returns a string representation of the numeral. *)
- val to_string : bitvec_num -> string
+ val to_string : Expr.expr -> string
(** Creates a bit-vector constant. *)
- val mk_const : context -> Symbol.symbol -> int -> bitvec_expr
+ val mk_const : context -> Symbol.symbol -> int -> Expr.expr
(** Creates a bit-vector constant. *)
- val mk_const_s : context -> string -> int -> bitvec_expr
+ val mk_const_s : context -> string -> int -> Expr.expr
(** Bitwise negation.
The argument must have a bit-vector sort. *)
- val mk_not : context -> bitvec_expr -> Expr.expr
+ val mk_not : context -> Expr.expr -> Expr.expr
(** Take conjunction of bits in a vector,vector of length 1.
The argument must have a bit-vector sort. *)
- val mk_redand : context -> bitvec_expr -> Expr.expr
+ val mk_redand : context -> Expr.expr -> Expr.expr
(** Take disjunction of bits in a vector,vector of length 1.
The argument must have a bit-vector sort. *)
- val mk_redor : context -> bitvec_expr -> Expr.expr
+ val mk_redor : context -> Expr.expr -> Expr.expr
(** Bitwise conjunction.
The arguments must have a bit-vector sort. *)
- val mk_and : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_and : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Bitwise disjunction.
The arguments must have a bit-vector sort. *)
- val mk_or : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_or : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Bitwise XOR.
The arguments must have a bit-vector sort. *)
- val mk_xor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_xor : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Bitwise NAND.
The arguments must have a bit-vector sort. *)
- val mk_nand : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_nand : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Bitwise NOR.
The arguments must have a bit-vector sort. *)
- val mk_nor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_nor : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Bitwise XNOR.
The arguments must have a bit-vector sort. *)
- val mk_xnor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_xnor : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Standard two's complement unary minus.
The arguments must have a bit-vector sort. *)
- val mk_neg : context -> bitvec_expr -> bitvec_expr
+ val mk_neg : context -> Expr.expr -> Expr.expr
(** Two's complement addition.
The arguments must have the same bit-vector sort. *)
- val mk_add : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_add : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Two's complement subtraction.
The arguments must have the same bit-vector sort. *)
- val mk_sub : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_sub : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Two's complement multiplication.
The arguments must have the same bit-vector sort. *)
- val mk_mul : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_mul : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Unsigned division.
@@ -1697,7 +1587,7 @@ sig
different from zero. If t2 is zero, then the result
is undefined.
The arguments must have the same bit-vector sort. *)
- val mk_udiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_udiv : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Signed division.
@@ -1709,14 +1599,14 @@ sig
If t2 is zero, then the result is undefined.
The arguments must have the same bit-vector sort. *)
- val mk_sdiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_sdiv : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Unsigned remainder.
It is defined as t1 - (t1 /u t2) * t2, where /u represents unsigned division.
If t2 is zero, then the result is undefined.
The arguments must have the same bit-vector sort. *)
- val mk_urem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_urem : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Signed remainder.
@@ -1725,53 +1615,53 @@ sig
If t2 is zero, then the result is undefined.
The arguments must have the same bit-vector sort. *)
- val mk_srem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_srem : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Two's complement signed remainder (sign follows divisor).
If t2 is zero, then the result is undefined.
The arguments must have the same bit-vector sort. *)
- val mk_smod : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_smod : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Unsigned less-than
The arguments must have the same bit-vector sort. *)
- val mk_ult : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_ult : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Two's complement signed less-than
The arguments must have the same bit-vector sort. *)
- val mk_slt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_slt : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Unsigned less-than or equal to.
The arguments must have the same bit-vector sort. *)
- val mk_ule : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_ule : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Two's complement signed less-than or equal to.
The arguments must have the same bit-vector sort. *)
- val mk_sle : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sle : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Unsigned greater than or equal to.
The arguments must have the same bit-vector sort. *)
- val mk_uge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_uge : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Two's complement signed greater than or equal to.
The arguments must have the same bit-vector sort. *)
- val mk_sge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sge : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Unsigned greater-than.
The arguments must have the same bit-vector sort. *)
- val mk_ugt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_ugt : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Two's complement signed greater-than.
The arguments must have the same bit-vector sort. *)
- val mk_sgt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sgt : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Bit-vector concatenation.
@@ -1779,30 +1669,30 @@ sig
@return
The result is a bit-vector of size n1+n2, where n1 (n2)
is the size of t1 (t2). *)
- val mk_concat : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_concat : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Bit-vector extraction.
Extract the bits between two limits from a bitvector of
size m to yield a new bitvector of size n, where
n = high - low + 1. *)
- val mk_extract : context -> int -> int -> bitvec_expr -> bitvec_expr
+ val mk_extract : context -> int -> int -> Expr.expr -> Expr.expr
(** Bit-vector sign extension.
Sign-extends the given bit-vector to the (signed) equivalent bitvector of
size m+i, where \c m is the size of the given bit-vector. *)
- val mk_sign_ext : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_sign_ext : context -> int -> Expr.expr -> Expr.expr
(** Bit-vector zero extension.
Extend the given bit-vector with zeros to the (unsigned) equivalent
bitvector of size m+i, where \c m is the size of the
given bit-vector. *)
- val mk_zero_ext : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_zero_ext : context -> int -> Expr.expr -> Expr.expr
(** Bit-vector repetition. *)
- val mk_repeat : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_repeat : context -> int -> Expr.expr -> Expr.expr
(** Shift left.
@@ -1811,7 +1701,7 @@ sig
NB. The semantics of shift operations varies between environments. This
definition does not necessarily capture directly the semantics of the
programming language or assembly architecture you are modeling.*)
- val mk_shl : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_shl : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Logical shift right
@@ -1822,7 +1712,7 @@ sig
programming language or assembly architecture you are modeling.
The arguments must have a bit-vector sort. *)
- val mk_lshr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_lshr : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Arithmetic shift right
@@ -1835,23 +1725,23 @@ sig
programming language or assembly architecture you are modeling.
The arguments must have a bit-vector sort. *)
- val mk_ashr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_ashr : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Rotate Left.
Rotate bits of \c t to the left \c i times. *)
- val mk_rotate_left : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_rotate_left : context -> int -> Expr.expr -> Expr.expr
(** Rotate Right.
Rotate bits of \c t to the right \c i times.*)
- val mk_rotate_right : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_rotate_right : context -> int -> Expr.expr -> Expr.expr
(** Rotate Left.
Rotate bits of the second argument to the left.*)
- val mk_ext_rotate_left : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_ext_rotate_left : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Rotate Right.
Rotate bits of the second argument to the right. *)
- val mk_ext_rotate_right : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_ext_rotate_right : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create an integer from the bit-vector argument
@@ -1863,50 +1753,50 @@ sig
NB. This function is essentially treated as uninterpreted.
So you cannot expect Z3 to precisely reflect the semantics of this function
when solving constraints with this function.*)
- val mk_bv2int : context -> bitvec_expr -> bool -> Arithmetic.Integer.int_expr
+ val mk_bv2int : context -> Expr.expr -> bool -> Expr.expr
(** Create a predicate that checks that the bit-wise addition does not overflow.
The arguments must be of bit-vector sort. *)
- val mk_add_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+ val mk_add_no_overflow : context -> Expr.expr -> Expr.expr -> bool -> Expr.expr
(** Create a predicate that checks that the bit-wise addition does not underflow.
The arguments must be of bit-vector sort. *)
- val mk_add_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_add_no_underflow : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create a predicate that checks that the bit-wise subtraction does not overflow.
The arguments must be of bit-vector sort. *)
- val mk_sub_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sub_no_overflow : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create a predicate that checks that the bit-wise subtraction does not underflow.
The arguments must be of bit-vector sort. *)
- val mk_sub_no_underflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+ val mk_sub_no_underflow : context -> Expr.expr -> Expr.expr -> bool -> Expr.expr
(** Create a predicate that checks that the bit-wise signed division does not overflow.
The arguments must be of bit-vector sort. *)
- val mk_sdiv_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sdiv_no_overflow : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create a predicate that checks that the bit-wise negation does not overflow.
The arguments must be of bit-vector sort. *)
- val mk_neg_no_overflow : context -> bitvec_expr -> Boolean.bool_expr
+ val mk_neg_no_overflow : context -> Expr.expr -> Expr.expr
(** Create a predicate that checks that the bit-wise multiplication does not overflow.
The arguments must be of bit-vector sort. *)
- val mk_mul_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+ val mk_mul_no_overflow : context -> Expr.expr -> Expr.expr -> bool -> Expr.expr
(** Create a predicate that checks that the bit-wise multiplication does not underflow.
The arguments must be of bit-vector sort. *)
- val mk_mul_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_mul_no_underflow : context -> Expr.expr -> Expr.expr -> Expr.expr
(** Create a bit-vector numeral. *)
- val mk_numeral : context -> string -> int -> bitvec_num
+ val mk_numeral : context -> string -> int -> Expr.expr
end
(** Functions to manipulate proof expressions *)
@@ -2332,7 +2222,7 @@ sig
val is_garbage : goal -> bool
(** Adds the constraints to the given goal. *)
- val assert_ : goal -> Boolean.bool_expr list -> unit
+ val assert_ : goal -> Expr.expr list -> unit
(** Indicates whether the goal contains `false'. *)
val is_inconsistent : goal -> bool
@@ -2348,7 +2238,7 @@ sig
val get_size : goal -> int
(** The formulas in the goal. *)
- val get_formulas : goal -> Boolean.bool_expr list
+ val get_formulas : goal -> Expr.expr list
(** The number of formulas, subformulas and terms in the goal. *)
val get_num_exprs : goal -> int
@@ -2759,7 +2649,7 @@ sig
val reset : solver -> unit
(** Assert a constraint (or multiple) into the solver. *)
- val assert_ : solver -> Boolean.bool_expr list -> unit
+ val assert_ : solver -> Expr.expr list -> unit
(** * Assert multiple constraints (cs) into the solver, and track them (in the
* unsat) core
@@ -2772,7 +2662,7 @@ sig
* of the Boolean variables provided using {!assert_and_track}
* and the Boolean literals
* provided using {!check} with assumptions. *)
- val assert_and_track_a : solver -> Boolean.bool_expr list -> Boolean.bool_expr list -> unit
+ val assert_and_track_a : solver -> Expr.expr list -> Expr.expr list -> unit
(** * Assert a constraint (c) into the solver, and track it (in the unsat) core
* using the Boolean constant p.
@@ -2784,20 +2674,20 @@ sig
* of the Boolean variables provided using {!assert_and_track}
* and the Boolean literals
* provided using {!check} with assumptions. *)
- val assert_and_track : solver -> Boolean.bool_expr -> Boolean.bool_expr -> unit
+ val assert_and_track : solver -> Expr.expr -> Expr.expr -> unit
(** The number of assertions in the solver. *)
val get_num_assertions : solver -> int
(** The set of asserted formulas. *)
- val get_assertions : solver -> Boolean.bool_expr list
+ val get_assertions : solver -> Expr.expr list
(** Checks whether the assertions in the solver are consistent or not.
{!Model}
{!get_unsat_core}
{!Proof} *)
- val check : solver -> Boolean.bool_expr list -> status
+ val check : solver -> Expr.expr list -> status
(** The model of the last Check.
@@ -2863,13 +2753,13 @@ sig
val get_param_descrs : fixedpoint -> Params.ParamDescrs.param_descrs
(** Assert a constraints into the fixedpoint solver. *)
- val assert_ : fixedpoint -> Boolean.bool_expr list -> unit
+ val assert_ : fixedpoint -> Expr.expr list -> unit
(** Register predicate as recursive relation. *)
val register_relation : fixedpoint -> FuncDecl.func_decl -> unit
(** Add rule into the fixedpoint solver. *)
- val add_rule : fixedpoint -> Boolean.bool_expr -> Symbol.symbol option -> unit
+ val add_rule : fixedpoint -> Expr.expr -> Symbol.symbol option -> unit
(** Add table fact to the fixedpoint solver. *)
val add_fact : fixedpoint -> FuncDecl.func_decl -> int list -> unit
@@ -2878,7 +2768,7 @@ sig
A query is a conjunction of constraints. The constraints may include the recursively defined relations.
The query is satisfiable if there is an instance of the query variables and a derivation for it.
The query is unsatisfiable if there are no derivations satisfying the query variables. *)
- val query : fixedpoint -> Boolean.bool_expr -> Solver.status
+ val query : fixedpoint -> Expr.expr -> Solver.status
(** Query the fixedpoint solver.
A query is an array of relations.
@@ -2897,7 +2787,7 @@ sig
val pop : fixedpoint -> unit
(** Update named rule into in the fixedpoint solver. *)
- val update_rule : fixedpoint -> Boolean.bool_expr -> Symbol.symbol -> unit
+ val update_rule : fixedpoint -> Expr.expr -> Symbol.symbol -> unit
(** Retrieve satisfying instance or instances of solver,
or definitions for the recursive predicates that show unsatisfiability. *)
@@ -2923,13 +2813,13 @@ sig
val set_predicate_representation : fixedpoint -> FuncDecl.func_decl -> Symbol.symbol list -> unit
(** Convert benchmark given as set of axioms, rules and queries to a string. *)
- val to_string_q : fixedpoint -> Boolean.bool_expr list -> string
+ val to_string_q : fixedpoint -> Expr.expr list -> string
(** Retrieve set of rules added to fixedpoint context. *)
- val get_rules : fixedpoint -> Boolean.bool_expr list
+ val get_rules : fixedpoint -> Expr.expr list
(** Retrieve set of assertions added to fixedpoint context. *)
- val get_assertions : fixedpoint -> Boolean.bool_expr list
+ val get_assertions : fixedpoint -> Expr.expr list
(** Create a Fixedpoint context. *)
val mk_fixedpoint : context -> fixedpoint
@@ -2985,7 +2875,7 @@ sig
(** Convert a benchmark into an SMT-LIB formatted string.
@return A string representation of the benchmark. *)
- val benchmark_to_smtstring : context -> string -> string -> string -> string -> Boolean.bool_expr list -> Boolean.bool_expr -> string
+ val benchmark_to_smtstring : context -> string -> string -> string -> string -> Expr.expr list -> Expr.expr -> string
(** Parse the given string using the SMT-LIB parser.
@@ -3004,13 +2894,13 @@ sig
val get_num_smtlib_formulas : context -> int
(** The formulas parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
- val get_smtlib_formulas : context -> Boolean.bool_expr list
+ val get_smtlib_formulas : context -> Expr.expr list
(** The number of SMTLIB assumptions parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
val get_num_smtlib_assumptions : context -> int
(** The assumptions parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
- val get_smtlib_assumptions : context -> Boolean.bool_expr list
+ val get_smtlib_assumptions : context -> Expr.expr list
(** The number of SMTLIB declarations parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
val get_num_smtlib_decls : context -> int
@@ -3028,11 +2918,11 @@ sig
{!parse_smtlib_string}
@return A conjunction of assertions in the scope (up to push/pop) at the end of the string. *)
- val parse_smtlib2_string : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> Boolean.bool_expr
+ val parse_smtlib2_string : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> Expr.expr
(** Parse the given file using the SMT-LIB2 parser.
{!parse_smtlib2_string} *)
- val parse_smtlib2_file : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> Boolean.bool_expr
+ val parse_smtlib2_file : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> Expr.expr
end
(** Set a global (or module) parameter, which is shared by all Z3 contexts.
diff --git a/src/api/ml/z3_rich.ml b/src/api/ml/z3_rich.ml
new file mode 100644
index 000000000..052ca5e22
--- /dev/null
+++ b/src/api/ml/z3_rich.ml
@@ -0,0 +1,3400 @@
+(**
+ The Z3 ML/Ocaml Interface.
+
+ Copyright (C) 2012 Microsoft Corporation
+ @author CM Wintersteiger (cwinter) 2012-12-17
+*)
+
+open Z3enums
+
+(* Some helpers. *)
+let null = Z3native.mk_null()
+let is_null o = (Z3native.is_null o)
+
+
+(* Internal types *)
+type z3_native_context = { m_n_ctx : Z3native.z3_context; m_n_obj_cnt: int; }
+type context = z3_native_context
+
+type z3_native_object = {
+ m_ctx : context ;
+ mutable m_n_obj : Z3native.ptr ;
+ inc_ref : Z3native.z3_context -> Z3native.ptr -> unit;
+ dec_ref : Z3native.z3_context -> Z3native.ptr -> unit }
+
+
+(** Internal stuff *)
+module Internal =
+struct
+ let dispose_context ctx =
+ if ctx.m_n_obj_cnt == 0 then (
+ (Z3native.del_context ctx.m_n_ctx)
+ ) else (
+ Printf.printf "ERROR: NOT DISPOSING CONTEXT (because it still has %d objects alive)\n" ctx.m_n_obj_cnt;
+ )
+
+ let create_context settings =
+ let cfg = Z3native.mk_config in
+ let f e = (Z3native.set_param_value cfg (fst e) (snd e)) in
+ (List.iter f settings) ;
+ let v = Z3native.mk_context_rc cfg in
+ Z3native.del_config(cfg) ;
+ Z3native.set_ast_print_mode v (int_of_ast_print_mode PRINT_SMTLIB2_COMPLIANT) ;
+ Z3native.set_internal_error_handler v ;
+ let res = { m_n_ctx = v; m_n_obj_cnt = 0 } in
+ let f = fun o -> dispose_context o in
+ Gc.finalise f res;
+ res
+
+ let context_add1 ctx = ignore (ctx.m_n_obj_cnt = ctx.m_n_obj_cnt + 1)
+ let context_sub1 ctx = ignore (ctx.m_n_obj_cnt = ctx.m_n_obj_cnt - 1)
+ let context_gno ctx = ctx.m_n_ctx
+
+
+ let z3obj_gc o = o.m_ctx
+ let z3obj_gnc o = (context_gno o.m_ctx)
+
+ let z3obj_gno o = o.m_n_obj
+ let z3obj_sno o ctx no =
+ (context_add1 ctx) ;
+ o.inc_ref (context_gno ctx) no ;
+ (
+ if not (is_null o.m_n_obj) then
+ o.dec_ref (context_gno ctx) o.m_n_obj ;
+ (context_sub1 ctx)
+ ) ;
+ o.m_n_obj <- no
+
+ let z3obj_dispose o =
+ if not (is_null o.m_n_obj) then
+ (
+ o.dec_ref (z3obj_gnc o) o.m_n_obj ;
+ (context_sub1 (z3obj_gc o))
+ ) ;
+ o.m_n_obj <- null
+
+ let z3obj_create o =
+ let f = fun o -> (z3obj_dispose o) in
+ Gc.finalise f o
+
+ let z3obj_nil_ref x y = ()
+
+ let z3_native_object_of_ast_ptr : context -> Z3native.ptr -> z3_native_object = fun ctx no ->
+ let res : z3_native_object = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.inc_ref ;
+ dec_ref = Z3native.dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+end
+
+open Internal
+
+module Log =
+struct
+ let open_ filename = ((lbool_of_int (Z3native.open_log filename)) == L_TRUE)
+ let close = Z3native.close_log
+ let append s = Z3native.append_log s
+end
+
+
+module Version =
+struct
+ let major = let (x, _, _, _) = Z3native.get_version in x
+ let minor = let (_, x, _, _) = Z3native.get_version in x
+ let build = let (_, _, x, _) = Z3native.get_version in x
+ let revision = let (_, _, _, x) = Z3native.get_version in x
+ let to_string =
+ let (mj, mn, bld, rev) = Z3native.get_version in
+ string_of_int mj ^ "." ^
+ string_of_int mn ^ "." ^
+ string_of_int bld ^ "." ^
+ string_of_int rev ^ "."
+end
+
+
+let mk_list ( f : int -> 'a ) ( n : int ) =
+ let rec mk_list' ( f : int -> 'a ) ( i : int ) ( n : int ) ( tail : 'a list ) : 'a list =
+ if (i >= n) then
+ tail
+ else
+ (mk_list' f (i+1) n ((f i) :: tail))
+ in
+ mk_list' f 0 n []
+
+let list_of_array ( x : _ array ) =
+ let f i = (Array.get x i) in
+ mk_list f (Array.length x)
+
+let mk_context ( cfg : ( string * string ) list ) =
+ create_context cfg
+
+
+module Symbol =
+struct
+ (* Symbol types *)
+ type int_symbol = z3_native_object
+ type string_symbol = z3_native_object
+
+ type symbol =
+ | S_Int of int_symbol
+ | S_Str of string_symbol
+
+
+ let create_i ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : int_symbol = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = z3obj_nil_ref ;
+ dec_ref = z3obj_nil_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let create_s ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : string_symbol = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = z3obj_nil_ref ;
+ dec_ref = z3obj_nil_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ match (symbol_kind_of_int (Z3native.get_symbol_kind (context_gno ctx) no)) with
+ | INT_SYMBOL -> S_Int (create_i ctx no)
+ | STRING_SYMBOL -> S_Str (create_s ctx no)
+
+ let gc ( x : symbol ) =
+ match x with
+ | S_Int(n) -> (z3obj_gc n)
+ | S_Str(n) -> (z3obj_gc n)
+
+ let gnc ( x : symbol ) =
+ match x with
+ | S_Int(n) -> (z3obj_gnc n)
+ | S_Str(n) -> (z3obj_gnc n)
+
+ let gno ( x : symbol ) =
+ match x with
+ | S_Int(n) -> (z3obj_gno n)
+ | S_Str(n) -> (z3obj_gno n)
+
+ let symbol_lton ( a : symbol list ) =
+ let f ( e : symbol ) = (gno e) in
+ Array.of_list (List.map f a)
+
+ let kind ( o : symbol ) = (symbol_kind_of_int (Z3native.get_symbol_kind (gnc o) (gno o)))
+ let is_int_symbol ( o : symbol ) = (kind o) == INT_SYMBOL
+ let is_string_symbol ( o : symbol ) = (kind o) == STRING_SYMBOL
+ let get_int (o : int_symbol) = Z3native.get_symbol_int (z3obj_gnc o) (z3obj_gno o)
+ let get_string (o : string_symbol) = Z3native.get_symbol_string (z3obj_gnc o) (z3obj_gno o)
+ let to_string ( o : symbol ) =
+ match (kind o) with
+ | INT_SYMBOL -> (string_of_int (Z3native.get_symbol_int (gnc o) (gno o)))
+ | STRING_SYMBOL -> (Z3native.get_symbol_string (gnc o) (gno o))
+
+ let mk_int ( ctx : context ) ( i : int ) =
+ S_Int (create_i ctx (Z3native.mk_int_symbol (context_gno ctx) i))
+
+ let mk_string ( ctx : context ) ( s : string ) =
+ S_Str (create_s ctx (Z3native.mk_string_symbol (context_gno ctx) s))
+
+ let mk_ints ( ctx : context ) ( names : int list ) =
+ let f elem = mk_int ( ctx : context ) elem in
+ (List.map f names)
+
+ let mk_strings ( ctx : context ) ( names : string list ) =
+ let f elem = mk_string ( ctx : context ) elem in
+ (List.map f names)
+end
+
+
+module AST =
+struct
+ type ast = z3_native_object
+
+ let context_of_ast ( x : ast ) = (z3obj_gc x)
+ let nc_of_ast ( x : ast ) = (z3obj_gnc x)
+ let ptr_of_ast ( x : ast ) = (z3obj_gno x)
+
+ let rec ast_of_ptr : context -> Z3native.ptr -> ast = fun ctx no ->
+ match (ast_kind_of_int (Z3native.get_ast_kind (context_gno ctx) no)) with
+ | FUNC_DECL_AST
+ | SORT_AST
+ | QUANTIFIER_AST
+ | APP_AST
+ | NUMERAL_AST
+ | VAR_AST -> z3_native_object_of_ast_ptr ctx no
+ | UNKNOWN_AST -> raise (Z3native.Exception "Cannot create asts of type unknown")
+
+ module ASTVector =
+ struct
+ type ast_vector = z3_native_object
+
+ let ast_vector_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : ast_vector = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.ast_vector_inc_ref ;
+ dec_ref = Z3native.ast_vector_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let get_size ( x : ast_vector ) =
+ Z3native.ast_vector_size (z3obj_gnc x) (z3obj_gno x)
+
+ let get ( x : ast_vector ) ( i : int ) =
+ ast_of_ptr (z3obj_gc x) (Z3native.ast_vector_get (z3obj_gnc x) (z3obj_gno x) i)
+
+ let set ( x : ast_vector ) ( i : int ) ( value : ast ) =
+ Z3native.ast_vector_set (z3obj_gnc x) (z3obj_gno x) i (z3obj_gno value)
+
+ let resize ( x : ast_vector ) ( new_size : int ) =
+ Z3native.ast_vector_resize (z3obj_gnc x) (z3obj_gno x) new_size
+
+ let push ( x : ast_vector ) ( a : ast ) =
+ Z3native.ast_vector_push (z3obj_gnc x) (z3obj_gno x) (z3obj_gno a)
+
+ let translate ( x : ast_vector ) ( to_ctx : context ) =
+ ast_vector_of_ptr to_ctx (Z3native.ast_vector_translate (z3obj_gnc x) (z3obj_gno x) (context_gno to_ctx))
+
+ let to_string ( x : ast_vector ) =
+ Z3native.ast_vector_to_string (z3obj_gnc x) (z3obj_gno x)
+ end
+
+ module ASTMap =
+ struct
+ type ast_map = z3_native_object
+
+ let astmap_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : ast_map = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.ast_map_inc_ref ;
+ dec_ref = Z3native.ast_map_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let contains ( x : ast_map ) ( key : ast ) =
+ Z3native.ast_map_contains (z3obj_gnc x) (z3obj_gno x) (z3obj_gno key)
+
+ let find ( x : ast_map ) ( key : ast ) =
+ ast_of_ptr (z3obj_gc x) (Z3native.ast_map_find (z3obj_gnc x) (z3obj_gno x) (z3obj_gno key))
+
+ let insert ( x : ast_map ) ( key : ast ) ( value : ast) =
+ Z3native.ast_map_insert (z3obj_gnc x) (z3obj_gno x) (z3obj_gno key) (z3obj_gno value)
+
+ let erase ( x : ast_map ) ( key : ast ) =
+ Z3native.ast_map_erase (z3obj_gnc x) (z3obj_gno x) (z3obj_gno key)
+
+ let reset ( x : ast_map ) =
+ Z3native.ast_map_reset (z3obj_gnc x) (z3obj_gno x)
+
+ let get_size ( x : ast_map ) =
+ Z3native.ast_map_size (z3obj_gnc x) (z3obj_gno x)
+
+ let get_keys ( x : ast_map ) =
+ let av = ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.ast_map_keys (z3obj_gnc x) (z3obj_gno x)) in
+ let f i = (ASTVector.get av i) in
+ mk_list f (ASTVector.get_size av)
+
+ let to_string ( x : ast_map ) =
+ Z3native.ast_map_to_string (z3obj_gnc x) (z3obj_gno x)
+ end
+
+ let get_hash_code ( x : ast ) = Z3native.get_ast_hash (z3obj_gnc x) (z3obj_gno x)
+ let get_id ( x : ast ) = Z3native.get_ast_id (z3obj_gnc x) (z3obj_gno x)
+ let get_ast_kind ( x : ast ) = (ast_kind_of_int (Z3native.get_ast_kind (z3obj_gnc x) (z3obj_gno x)))
+
+ let is_expr ( x : ast ) =
+ match get_ast_kind ( x : ast ) with
+ | APP_AST
+ | NUMERAL_AST
+ | QUANTIFIER_AST
+ | VAR_AST -> true
+ | _ -> false
+
+ let is_var ( x : ast ) = (get_ast_kind x) == VAR_AST
+ let is_quantifier ( x : ast ) = (get_ast_kind x) == QUANTIFIER_AST
+ let is_sort ( x : ast ) = (get_ast_kind x) == SORT_AST
+ let is_func_decl ( x : ast ) = (get_ast_kind x) == FUNC_DECL_AST
+
+ let to_string ( x : ast ) = Z3native.ast_to_string (z3obj_gnc x) (z3obj_gno x)
+ let to_sexpr ( x : ast ) = Z3native.ast_to_string (z3obj_gnc x) (z3obj_gno x)
+
+
+ let ( = ) ( a : ast ) ( b : ast ) = (a == b) ||
+ if (z3obj_gnc a) != (z3obj_gnc b) then
+ false
+ else
+ Z3native.is_eq_ast (z3obj_gnc a) (z3obj_gno a) (z3obj_gno b)
+
+ let compare a b =
+ if (get_id a) < (get_id b) then -1 else
+ if (get_id a) > (get_id b) then 1 else
+ 0
+
+ let ( < ) (a : ast) (b : ast) = (compare a b)
+
+ let translate ( x : ast ) ( to_ctx : context ) =
+ if (z3obj_gnc x) == (context_gno to_ctx) then
+ x
+ else
+ ast_of_ptr to_ctx (Z3native.translate (z3obj_gnc x) (z3obj_gno x) (context_gno to_ctx))
+
+ let wrap_ast ( ctx : context ) ( ptr : Z3native.ptr ) = ast_of_ptr ctx ptr
+ let unwrap_ast ( x : ast ) = (z3obj_gno x)
+end
+
+open AST
+
+
+module Sort =
+struct
+ type sort = Sort of AST.ast
+ type uninterpreted_sort = UninterpretedSort of sort
+
+ let sort_of_ptr : context -> Z3native.ptr -> sort = fun ctx no ->
+ let q = (z3_native_object_of_ast_ptr ctx no) in
+ if ((Z3enums.ast_kind_of_int (Z3native.get_ast_kind (context_gno ctx) no)) != Z3enums.SORT_AST) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ match (sort_kind_of_int (Z3native.get_sort_kind (context_gno ctx) no)) with
+ | ARRAY_SORT
+ | BOOL_SORT
+ | BV_SORT
+ | DATATYPE_SORT
+ | INT_SORT
+ | REAL_SORT
+ | UNINTERPRETED_SORT
+ | FINITE_DOMAIN_SORT
+ | RELATION_SORT -> Sort(q)
+ | UNKNOWN_SORT -> raise (Z3native.Exception "Unknown sort kind encountered")
+
+ let ast_of_sort s = match s with Sort(x) -> x
+ let sort_of_uninterpreted_sort s = match s with UninterpretedSort(x) -> x
+
+ let uninterpreted_sort_of_sort s = match s with Sort(a) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.UNINTERPRETED_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ UninterpretedSort(s)
+
+ let gc ( x : sort ) = (match x with Sort(a) -> (z3obj_gc a))
+ let gnc ( x : sort ) = (match x with Sort(a) -> (z3obj_gnc a))
+ let gno ( x : sort ) = (match x with Sort(a) -> (z3obj_gno a))
+
+ let sort_lton ( a : sort list ) =
+ let f ( e : sort ) = match e with Sort(a) -> (AST.ptr_of_ast a) in
+ Array.of_list (List.map f a)
+
+ let ( = ) : sort -> sort -> bool = fun a b ->
+ (a == b) ||
+ if (gnc a) != (gnc b) then
+ false
+ else
+ (Z3native.is_eq_sort (gnc a) (gno a) (gno b))
+
+
+ let get_id ( x : sort ) = Z3native.get_sort_id (gnc x) (gno x)
+ let get_sort_kind ( x : sort ) = (sort_kind_of_int (Z3native.get_sort_kind (gnc x) (gno x)))
+ let get_name ( x : sort ) = (Symbol.create (gc x) (Z3native.get_sort_name (gnc x) (gno x)))
+ let to_string ( x : sort ) = Z3native.sort_to_string (gnc x) (gno x)
+
+ let mk_uninterpreted ( ctx : context ) ( s : Symbol.symbol ) =
+ let res = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.inc_ref ;
+ dec_ref = Z3native.dec_ref } in
+ (z3obj_sno res ctx (Z3native.mk_uninterpreted_sort (context_gno ctx) (Symbol.gno s))) ;
+ (z3obj_create res) ;
+ UninterpretedSort(Sort(res))
+
+ let mk_uninterpreted_s ( ctx : context ) ( s : string ) =
+ mk_uninterpreted ctx (Symbol.mk_string ( ctx : context ) s)
+end
+
+open Sort
+
+
+module rec FuncDecl :
+sig
+ type func_decl = FuncDecl of AST.ast
+ val ast_of_func_decl : FuncDecl.func_decl -> AST.ast
+ val func_decl_of_ptr : context -> Z3native.ptr -> func_decl
+ val gc : func_decl -> context
+ val gnc : func_decl -> Z3native.ptr
+ val gno : func_decl -> Z3native.ptr
+ module Parameter :
+ sig
+ type parameter =
+ P_Int of int
+ | P_Dbl of float
+ | P_Sym of Symbol.symbol
+ | P_Srt of Sort.sort
+ | P_Ast of AST.ast
+ | P_Fdl of func_decl
+ | P_Rat of string
+
+ val get_kind : parameter -> Z3enums.parameter_kind
+ val get_int : parameter -> int
+ val get_float : parameter -> float
+ val get_symbol : parameter -> Symbol.symbol
+ val get_sort : parameter -> Sort.sort
+ val get_ast : parameter -> AST.ast
+ val get_func_decl : parameter -> func_decl
+ val get_rational : parameter -> string
+ end
+ val mk_func_decl : context -> Symbol.symbol -> Sort.sort list -> Sort.sort -> func_decl
+ val mk_func_decl_s : context -> string -> Sort.sort list -> Sort.sort -> func_decl
+ val mk_fresh_func_decl : context -> string -> Sort.sort list -> Sort.sort -> func_decl
+ val mk_const_decl : context -> Symbol.symbol -> Sort.sort -> func_decl
+ val mk_const_decl_s : context -> string -> Sort.sort -> func_decl
+ val mk_fresh_const_decl : context -> string -> Sort.sort -> func_decl
+ val ( = ) : func_decl -> func_decl -> bool
+ val to_string : func_decl -> string
+ val get_id : func_decl -> int
+ val get_arity : func_decl -> int
+ val get_domain_size : func_decl -> int
+ val get_domain : func_decl -> Sort.sort list
+ val get_range : func_decl -> Sort.sort
+ val get_decl_kind : func_decl -> Z3enums.decl_kind
+ val get_name : func_decl -> Symbol.symbol
+ val get_num_parameters : func_decl -> int
+ val get_parameters : func_decl -> Parameter.parameter list
+ val apply : func_decl -> Expr.expr list -> Expr.expr
+end = struct
+ type func_decl = FuncDecl of AST.ast
+
+ let func_decl_of_ptr : context -> Z3native.ptr -> func_decl = fun ctx no ->
+ if ((Z3enums.ast_kind_of_int (Z3native.get_ast_kind (context_gno ctx) no)) != Z3enums.FUNC_DECL_AST) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ FuncDecl(z3_native_object_of_ast_ptr ctx no)
+
+ let ast_of_func_decl f = match f with FuncDecl(x) -> x
+
+ let create_ndr ( ctx : context ) ( name : Symbol.symbol ) ( domain : sort list ) ( range : sort ) =
+ let res = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.inc_ref ;
+ dec_ref = Z3native.dec_ref } in
+ (z3obj_sno res ctx (Z3native.mk_func_decl (context_gno ctx) (Symbol.gno name) (List.length domain) (sort_lton domain) (Sort.gno range))) ;
+ (z3obj_create res) ;
+ FuncDecl(res)
+
+ let create_pdr ( ctx : context) ( prefix : string ) ( domain : sort list ) ( range : sort ) =
+ let res = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.inc_ref ;
+ dec_ref = Z3native.dec_ref } in
+ (z3obj_sno res ctx (Z3native.mk_fresh_func_decl (context_gno ctx) prefix (List.length domain) (sort_lton domain) (Sort.gno range))) ;
+ (z3obj_create res) ;
+ FuncDecl(res)
+
+ let gc ( x : func_decl ) = match x with FuncDecl(a) -> (z3obj_gc a)
+ let gnc ( x : func_decl ) = match x with FuncDecl(a) -> (z3obj_gnc a)
+ let gno ( x : func_decl ) = match x with FuncDecl(a) -> (z3obj_gno a)
+
+ module Parameter =
+ struct
+ type parameter =
+ | P_Int of int
+ | P_Dbl of float
+ | P_Sym of Symbol.symbol
+ | P_Srt of Sort.sort
+ | P_Ast of AST.ast
+ | P_Fdl of func_decl
+ | P_Rat of string
+
+ let get_kind ( x : parameter ) =
+ (match x with
+ | P_Int(_) -> PARAMETER_INT
+ | P_Dbl(_) -> PARAMETER_DOUBLE
+ | P_Sym(_) -> PARAMETER_SYMBOL
+ | P_Srt(_) -> PARAMETER_SORT
+ | P_Ast(_) -> PARAMETER_AST
+ | P_Fdl(_) -> PARAMETER_FUNC_DECL
+ | P_Rat(_) -> PARAMETER_RATIONAL)
+
+ let get_int ( x : parameter ) =
+ match x with
+ | P_Int(x) -> x
+ | _ -> raise (Z3native.Exception "parameter is not an int")
+
+ let get_float ( x : parameter ) =
+ match x with
+ | P_Dbl(x) -> x
+ | _ -> raise (Z3native.Exception "parameter is not a double")
+
+ let get_symbol ( x : parameter ) =
+ match x with
+ | P_Sym(x) -> x
+ | _ -> raise (Z3native.Exception "parameter is not a symbol")
+
+ let get_sort ( x : parameter ) =
+ match x with
+ | P_Srt(x) -> x
+ | _ -> raise (Z3native.Exception "parameter is not a sort")
+
+ let get_ast ( x : parameter ) =
+ match x with
+ | P_Ast(x) -> x
+ | _ -> raise (Z3native.Exception "parameter is not an ast")
+
+ let get_func_decl ( x : parameter ) =
+ match x with
+ | P_Fdl(x) -> x
+ | _ -> raise (Z3native.Exception "parameter is not a func_decl")
+
+ let get_rational ( x : parameter ) =
+ match x with
+ | P_Rat(x) -> x
+ | _ -> raise (Z3native.Exception "parameter is not a rational string")
+ end
+
+ let mk_func_decl ( ctx : context ) ( name : Symbol.symbol ) ( domain : sort list ) ( range : sort ) =
+ create_ndr ctx name domain range
+
+ let mk_func_decl_s ( ctx : context ) ( name : string ) ( domain : sort list ) ( range : sort ) =
+ mk_func_decl ctx (Symbol.mk_string ctx name) domain range
+
+ let mk_fresh_func_decl ( ctx : context ) ( prefix : string ) ( domain : sort list ) ( range : sort ) =
+ create_pdr ctx prefix domain range
+
+ let mk_const_decl ( ctx : context ) ( name : Symbol.symbol ) ( range : sort ) =
+ create_ndr ctx name [] range
+
+ let mk_const_decl_s ( ctx : context ) ( name : string ) ( range : sort ) =
+ create_ndr ctx (Symbol.mk_string ctx name) [] range
+
+ let mk_fresh_const_decl ( ctx : context ) ( prefix : string ) ( range : sort ) =
+ create_pdr ctx prefix [] range
+
+
+ let ( = ) ( a : func_decl ) ( b : func_decl ) = (a == b) ||
+ if (gnc a) != (gnc b) then
+ false
+ else
+ (Z3native.is_eq_func_decl (gnc a) (gno a) (gno b))
+
+ let to_string ( x : func_decl ) = Z3native.func_decl_to_string (gnc x) (gno x)
+
+ let get_id ( x : func_decl ) = Z3native.get_func_decl_id (gnc x) (gno x)
+
+ let get_arity ( x : func_decl ) = Z3native.get_arity (gnc x) (gno x)
+
+ let get_domain_size ( x : func_decl ) = Z3native.get_domain_size (gnc x) (gno x)
+
+ let get_domain ( x : func_decl ) =
+ let n = (get_domain_size x) in
+ let f i = sort_of_ptr (gc x) (Z3native.get_domain (gnc x) (gno x) i) in
+ mk_list f n
+
+ let get_range ( x : func_decl ) =
+ sort_of_ptr (gc x) (Z3native.get_range (gnc x) (gno x))
+
+ let get_decl_kind ( x : func_decl ) = (decl_kind_of_int (Z3native.get_decl_kind (gnc x) (gno x)))
+
+ let get_name ( x : func_decl ) = (Symbol.create (gc x) (Z3native.get_decl_name (gnc x) (gno x)))
+
+ let get_num_parameters ( x : func_decl ) = (Z3native.get_decl_num_parameters (gnc x) (gno x))
+
+ let get_parameters ( x : func_decl ) =
+ let n = (get_num_parameters x) in
+ let f i = (match (parameter_kind_of_int (Z3native.get_decl_parameter_kind (gnc x) (gno x) i)) with
+ | PARAMETER_INT -> Parameter.P_Int (Z3native.get_decl_int_parameter (gnc x) (gno x) i)
+ | PARAMETER_DOUBLE -> Parameter.P_Dbl (Z3native.get_decl_double_parameter (gnc x) (gno x) i)
+ | PARAMETER_SYMBOL-> Parameter.P_Sym (Symbol.create (gc x) (Z3native.get_decl_symbol_parameter (gnc x) (gno x) i))
+ | PARAMETER_SORT -> Parameter.P_Srt (sort_of_ptr (gc x) (Z3native.get_decl_sort_parameter (gnc x) (gno x) i))
+ | PARAMETER_AST -> Parameter.P_Ast (AST.ast_of_ptr (gc x) (Z3native.get_decl_ast_parameter (gnc x) (gno x) i))
+ | PARAMETER_FUNC_DECL -> Parameter.P_Fdl (func_decl_of_ptr (gc x) (Z3native.get_decl_func_decl_parameter (gnc x) (gno x) i))
+ | PARAMETER_RATIONAL -> Parameter.P_Rat (Z3native.get_decl_rational_parameter (gnc x) (gno x) i)
+ ) in
+ mk_list f n
+
+ let apply ( x : func_decl ) ( args : Expr.expr list ) = Expr.expr_of_func_app (gc x) x args
+end
+
+
+and Params :
+sig
+ type params = z3_native_object
+ module ParamDescrs :
+ sig
+ type param_descrs
+ val param_descrs_of_ptr : context -> Z3native.ptr -> param_descrs
+ val validate : param_descrs -> params -> unit
+ val get_kind : param_descrs -> Symbol.symbol -> Z3enums.param_kind
+ val get_names : param_descrs -> Symbol.symbol list
+ val get_size : param_descrs -> int
+ val to_string : param_descrs -> string
+ end
+ val add_bool : params -> Symbol.symbol -> bool -> unit
+ val add_int : params -> Symbol.symbol -> int -> unit
+ val add_double : params -> Symbol.symbol -> float -> unit
+ val add_symbol : params -> Symbol.symbol -> Symbol.symbol -> unit
+ val add_s_bool : params -> string -> bool -> unit
+ val add_s_int : params -> string -> int -> unit
+ val add_s_double : params -> string -> float -> unit
+ val add_s_symbol : params -> string -> Symbol.symbol -> unit
+ val mk_params : context -> params
+ val to_string : params -> string
+end = struct
+ type params = z3_native_object
+
+ module ParamDescrs =
+ struct
+ type param_descrs = z3_native_object
+
+ let param_descrs_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : param_descrs = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.param_descrs_inc_ref ;
+ dec_ref = Z3native.param_descrs_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let validate ( x : param_descrs ) ( p : params ) =
+ Z3native.params_validate (z3obj_gnc x) (z3obj_gno p) (z3obj_gno x)
+
+ let get_kind ( x : param_descrs ) ( name : Symbol.symbol ) =
+ (param_kind_of_int (Z3native.param_descrs_get_kind (z3obj_gnc x) (z3obj_gno x) (Symbol.gno name)))
+
+ let get_names ( x : param_descrs ) =
+ let n = Z3native.param_descrs_size (z3obj_gnc x) (z3obj_gno x) in
+ let f i = Symbol.create (z3obj_gc x) (Z3native.param_descrs_get_name (z3obj_gnc x) (z3obj_gno x) i) in
+ mk_list f n
+
+ let get_size ( x : param_descrs ) = Z3native.param_descrs_size (z3obj_gnc x) (z3obj_gno x)
+ let to_string ( x : param_descrs ) = Z3native.param_descrs_to_string (z3obj_gnc x) (z3obj_gno x)
+ end
+
+ let add_bool ( x : params ) ( name : Symbol.symbol ) ( value : bool ) =
+ Z3native.params_set_bool (z3obj_gnc x) (z3obj_gno x) (Symbol.gno name) value
+
+ let add_int ( x : params ) (name : Symbol.symbol ) ( value : int ) =
+ Z3native.params_set_uint (z3obj_gnc x) (z3obj_gno x) (Symbol.gno name) value
+
+ let add_double ( x : params ) ( name : Symbol.symbol ) ( value : float ) =
+ Z3native.params_set_double (z3obj_gnc x) (z3obj_gno x) (Symbol.gno name) value
+
+ let add_symbol ( x : params ) ( name : Symbol.symbol ) ( value : Symbol.symbol ) =
+ Z3native.params_set_symbol (z3obj_gnc x) (z3obj_gno x) (Symbol.gno name) (Symbol.gno value)
+
+ let add_s_bool ( x : params ) ( name : string ) ( value : bool ) =
+ add_bool x (Symbol.mk_string (z3obj_gc x) name) value
+
+ let add_s_int ( x : params) ( name : string ) ( value : int ) =
+ add_int x (Symbol.mk_string (z3obj_gc x) name) value
+
+ let add_s_double ( x : params ) ( name : string ) ( value : float ) =
+ add_double x (Symbol.mk_string (z3obj_gc x) name) value
+
+ let add_s_symbol ( x : params ) ( name : string ) ( value : Symbol.symbol ) =
+ add_symbol x (Symbol.mk_string (z3obj_gc x) name) value
+
+ let mk_params ( ctx : context ) =
+ let res : params = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.params_inc_ref ;
+ dec_ref = Z3native.params_dec_ref } in
+ (z3obj_sno res ctx (Z3native.mk_params (context_gno ctx))) ;
+ (z3obj_create res) ;
+ res
+
+ let to_string ( x : params ) = Z3native.params_to_string (z3obj_gnc x) (z3obj_gno x)
+end
+
+(** General expressions (terms) *)
+and Expr :
+sig
+ type expr = Expr of AST.ast
+ val expr_of_ptr : context -> Z3native.ptr -> expr
+ val gc : expr -> context
+ val gnc : expr -> Z3native.ptr
+ val gno : expr -> Z3native.ptr
+ val expr_lton : expr list -> Z3native.ptr array
+ val ast_of_expr : expr -> AST.ast
+ val expr_of_ast : AST.ast -> expr
+ val expr_of_func_app : context -> FuncDecl.func_decl -> expr list -> expr
+ val simplify : expr -> Params.params option -> expr
+ val get_simplify_help : context -> string
+ val get_simplify_parameter_descrs : context -> Params.ParamDescrs.param_descrs
+ val get_func_decl : expr -> FuncDecl.func_decl
+ val get_bool_value : expr -> Z3enums.lbool
+ val get_num_args : expr -> int
+ val get_args : expr -> expr list
+ val update : expr -> expr list -> expr
+ val substitute : expr -> expr list -> expr list -> expr
+ val substitute_one : expr -> expr -> expr -> expr
+ val substitute_vars : expr -> expr list -> expr
+ val translate : expr -> context -> expr
+ val to_string : expr -> string
+ val is_numeral : expr -> bool
+ val is_well_sorted : expr -> bool
+ val get_sort : expr -> Sort.sort
+ val is_bool : expr -> bool
+ val is_const : expr -> bool
+ val is_true : expr -> bool
+ val is_false : expr -> bool
+ val is_eq : expr -> bool
+ val is_distinct : expr -> bool
+ val is_ite : expr -> bool
+ val is_and : expr -> bool
+ val is_or : expr -> bool
+ val is_iff : expr -> bool
+ val is_xor : expr -> bool
+ val is_not : expr -> bool
+ val is_implies : expr -> bool
+ val is_label : expr -> bool
+ val is_label_lit : expr -> bool
+ val is_oeq : expr -> bool
+ val mk_const : context -> Symbol.symbol -> Sort.sort -> expr
+ val mk_const_s : context -> string -> Sort.sort -> expr
+ val mk_const_f : context -> FuncDecl.func_decl -> expr
+ val mk_fresh_const : context -> string -> Sort.sort -> expr
+ val mk_app : context -> FuncDecl.func_decl -> expr list -> expr
+ val mk_numeral_string : context -> string -> Sort.sort -> expr
+ val mk_numeral_int : context -> int -> Sort.sort -> expr
+end = struct
+ type expr = Expr of AST.ast
+
+ let c_of_expr e = match e with Expr(a) -> (z3obj_gc a)
+ let nc_of_expr e = match e with Expr(a) -> (z3obj_gnc a)
+ let ptr_of_expr e = match e with Expr(a) -> (z3obj_gno a)
+
+ let expr_of_ptr : context -> Z3native.ptr -> expr = fun ctx no ->
+ if ast_kind_of_int (Z3native.get_ast_kind (context_gno ctx) no) == QUANTIFIER_AST then
+ Expr(z3_native_object_of_ast_ptr ctx no)
+ else
+ let s = Z3native.get_sort (context_gno ctx) no in
+ let sk = (sort_kind_of_int (Z3native.get_sort_kind (context_gno ctx) s)) in
+ if (Z3native.is_algebraic_number (context_gno ctx) no) then
+ Expr(z3_native_object_of_ast_ptr ctx no)
+ else
+ if (Z3native.is_numeral_ast (context_gno ctx) no) then
+ if (sk == INT_SORT or sk == REAL_SORT or sk == BV_SORT) then
+ Expr(z3_native_object_of_ast_ptr ctx no)
+ else
+ raise (Z3native.Exception "Unsupported numeral object")
+ else
+ Expr(z3_native_object_of_ast_ptr ctx no)
+
+ let expr_of_ast a =
+ let q = (Z3enums.ast_kind_of_int (Z3native.get_ast_kind (z3obj_gnc a) (z3obj_gno a))) in
+ if (q != Z3enums.APP_AST && q != VAR_AST && q != QUANTIFIER_AST && q != NUMERAL_AST) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ Expr(a)
+
+ let ast_of_expr e = match e with Expr(a) -> a
+
+ let expr_lton ( a : expr list ) =
+ let f ( e : expr ) = match e with Expr(a) -> (AST.ptr_of_ast a) in
+ Array.of_list (List.map f a)
+
+ let expr_of_func_app : context -> FuncDecl.func_decl -> expr list -> expr = fun ctx f args ->
+ match f with FuncDecl.FuncDecl(fa) ->
+ let o = Z3native.mk_app (context_gno ctx) (AST.ptr_of_ast fa) (List.length args) (expr_lton args) in
+ expr_of_ptr ctx o
+
+ let simplify ( x : expr ) ( p : Params.params option ) = match p with
+ | None -> expr_of_ptr (Expr.gc x) (Z3native.simplify (Expr.gnc x) (Expr.gno x))
+ | Some pp -> expr_of_ptr (Expr.gc x) (Z3native.simplify_ex (Expr.gnc x) (Expr.gno x) (z3obj_gno pp))
+
+ let get_simplify_help ( ctx : context ) =
+ Z3native.simplify_get_help (context_gno ctx)
+
+ let get_simplify_parameter_descrs ( ctx : context ) =
+ Params.ParamDescrs.param_descrs_of_ptr ctx (Z3native.simplify_get_param_descrs (context_gno ctx))
+ let get_func_decl ( x : expr ) = FuncDecl.func_decl_of_ptr (Expr.gc x) (Z3native.get_app_decl (Expr.gnc x) (Expr.gno x))
+
+ let get_bool_value ( x : expr ) = lbool_of_int (Z3native.get_bool_value (Expr.gnc x) (Expr.gno x))
+
+ let get_num_args ( x : expr ) = Z3native.get_app_num_args (Expr.gnc x) (Expr.gno x)
+
+ let get_args ( x : expr ) = let n = (get_num_args x) in
+ let f i = expr_of_ptr (Expr.gc x) (Z3native.get_app_arg (Expr.gnc x) (Expr.gno x) i) in
+ mk_list f n
+
+ let update ( x : expr ) ( args : expr list ) =
+ if (List.length args <> (get_num_args x)) then
+ raise (Z3native.Exception "Number of arguments does not match")
+ else
+ expr_of_ptr (Expr.gc x) (Z3native.update_term (Expr.gnc x) (Expr.gno x) (List.length args) (expr_lton args))
+
+ let substitute ( x : expr ) from to_ =
+ if (List.length from) <> (List.length to_) then
+ raise (Z3native.Exception "Argument sizes do not match")
+ else
+ expr_of_ptr (Expr.gc x) (Z3native.substitute (Expr.gnc x) (Expr.gno x) (List.length from) (expr_lton from) (expr_lton to_))
+
+ let substitute_one ( x : expr ) from to_ =
+ substitute ( x : expr ) [ from ] [ to_ ]
+
+ let substitute_vars ( x : expr ) to_ =
+ expr_of_ptr (Expr.gc x) (Z3native.substitute_vars (Expr.gnc x) (Expr.gno x) (List.length to_) (expr_lton to_))
+
+ let translate ( x : expr ) to_ctx =
+ if (Expr.gc x) == to_ctx then
+ x
+ else
+ expr_of_ptr to_ctx (Z3native.translate (Expr.gnc x) (Expr.gno x) (context_gno to_ctx))
+
+ let to_string ( x : expr ) = Z3native.ast_to_string (Expr.gnc x) (Expr.gno x)
+
+ let is_numeral ( x : expr ) = (Z3native.is_numeral_ast (Expr.gnc x) (Expr.gno x))
+
+ let is_well_sorted ( x : expr ) = Z3native.is_well_sorted (Expr.gnc x) (Expr.gno x)
+
+ let get_sort ( x : expr ) = sort_of_ptr (Expr.gc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x))
+
+ let is_bool ( x : expr ) = (match x with Expr(a) -> (AST.is_expr a)) &&
+ (Z3native.is_eq_sort (Expr.gnc x)
+ (Z3native.mk_bool_sort (Expr.gnc x))
+ (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))
+
+ let is_const ( x : expr ) = (match x with Expr(a) -> (AST.is_expr a)) &&
+ (get_num_args x) == 0 &&
+ (FuncDecl.get_domain_size (get_func_decl x)) == 0
+
+ let is_true ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_TRUE)
+ let is_false ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_FALSE)
+ let is_eq ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_EQ)
+ let is_distinct ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_DISTINCT)
+ let is_ite ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_ITE)
+ let is_and ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_AND)
+ let is_or ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_OR)
+ let is_iff ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_IFF)
+ let is_xor ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_XOR)
+ let is_not ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_NOT)
+ let is_implies ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_IMPLIES)
+ let is_label ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_LABEL)
+ let is_label_lit ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_LABEL_LIT)
+ let is_oeq ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_OEQ)
+
+ let mk_const ( ctx : context ) ( name : Symbol.symbol ) ( range : sort ) =
+ expr_of_ptr ctx (Z3native.mk_const (context_gno ctx) (Symbol.gno name) (Sort.gno range))
+
+ let mk_const_s ( ctx : context ) ( name : string ) ( range : sort ) =
+ mk_const ctx (Symbol.mk_string ctx name) range
+
+ let mk_const_f ( ctx : context ) ( f : FuncDecl.func_decl ) = Expr.expr_of_func_app ctx f []
+
+ let mk_fresh_const ( ctx : context ) ( prefix : string ) ( range : sort ) =
+ expr_of_ptr ctx (Z3native.mk_fresh_const (context_gno ctx) prefix (Sort.gno range))
+
+ let mk_app ( ctx : context ) ( f : FuncDecl.func_decl ) ( args : expr list ) = expr_of_func_app ctx f args
+
+ let mk_numeral_string ( ctx : context ) ( v : string ) ( ty : sort ) =
+ expr_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (Sort.gno ty))
+
+ let mk_numeral_int ( ctx : context ) ( v : int ) ( ty : sort ) =
+ expr_of_ptr ctx (Z3native.mk_int (context_gno ctx) v (Sort.gno ty))
+end
+
+open FuncDecl
+open Expr
+
+module Boolean =
+struct
+ type bool_sort = BoolSort of Sort.sort
+ type bool_expr = BoolExpr of Expr.expr
+
+ let bool_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let a = (AST.ast_of_ptr ctx no) in
+ BoolExpr(Expr.Expr(a))
+
+ let bool_expr_of_expr e =
+ match e with Expr.Expr(a) ->
+ let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
+ if (q != Z3enums.BOOL_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ BoolExpr(e)
+
+ let bool_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ BoolSort(sort_of_ptr ctx no)
+
+ let sort_of_bool_sort s = match s with BoolSort(x) -> x
+
+ let bool_sort_of_sort s = match s with Sort(a) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.BOOL_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ BoolSort(s)
+
+ let expr_of_bool_expr e = match e with BoolExpr(x) -> x
+
+ let gc ( x : bool_expr ) = match x with BoolExpr(e) -> (Expr.c_of_expr e)
+ let gnc ( x : bool_expr ) = match x with BoolExpr(e) -> (Expr.nc_of_expr e)
+ let gno ( x : bool_expr ) = match x with BoolExpr(e) -> (Expr.ptr_of_expr e)
+
+ let mk_sort ( ctx : context ) =
+ BoolSort(sort_of_ptr ctx (Z3native.mk_bool_sort (context_gno ctx)))
+
+ let mk_const ( ctx : context ) ( name : Symbol.symbol ) =
+ let s = (match (mk_sort ctx) with BoolSort(q) -> q) in
+ BoolExpr(Expr.mk_const ctx name s)
+
+ let mk_const_s ( ctx : context ) ( name : string ) =
+ mk_const ctx (Symbol.mk_string ctx name)
+
+ let mk_true ( ctx : context ) =
+ bool_expr_of_ptr ctx (Z3native.mk_true (context_gno ctx))
+
+ let mk_false ( ctx : context ) =
+ bool_expr_of_ptr ctx (Z3native.mk_false (context_gno ctx))
+
+ let mk_val ( ctx : context ) ( value : bool ) =
+ if value then mk_true ctx else mk_false ctx
+
+ let mk_eq ( ctx : context ) ( x : expr ) ( y : expr ) =
+ bool_expr_of_ptr ctx (Z3native.mk_eq (context_gno ctx) (Expr.gno x) (Expr.gno y))
+
+ let mk_distinct ( ctx : context ) ( args : expr list ) =
+ bool_expr_of_ptr ctx (Z3native.mk_distinct (context_gno ctx) (List.length args) (expr_lton args))
+
+ let mk_not ( ctx : context ) ( a : bool_expr ) =
+ bool_expr_of_ptr ctx (Z3native.mk_not (context_gno ctx) (gno a))
+
+ let mk_ite ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) ( t3 : bool_expr ) =
+ bool_expr_of_ptr ctx (Z3native.mk_ite (context_gno ctx) (gno t1) (gno t2) (gno t3))
+
+ let mk_iff ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) =
+ bool_expr_of_ptr ctx (Z3native.mk_iff (context_gno ctx) (gno t1) (gno t2))
+
+ let mk_implies ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) =
+ bool_expr_of_ptr ctx (Z3native.mk_implies (context_gno ctx) (gno t1) (gno t2))
+
+ let mk_xor ( ctx : context ) ( t1 : bool_expr ) ( t2 : bool_expr ) =
+ bool_expr_of_ptr ctx (Z3native.mk_xor (context_gno ctx) (gno t1) (gno t2))
+
+ let mk_and ( ctx : context ) ( args : bool_expr list ) =
+ let f x = (Expr.gno (expr_of_bool_expr x)) in
+ bool_expr_of_ptr ctx (Z3native.mk_and (context_gno ctx) (List.length args) (Array.of_list (List.map f args)))
+
+ let mk_or ( ctx : context ) ( args : bool_expr list ) =
+ let f x = (Expr.gno (expr_of_bool_expr x)) in
+ bool_expr_of_ptr ctx (Z3native.mk_or (context_gno ctx) (List.length args) (Array.of_list(List.map f args)))
+end
+
+
+module Quantifier =
+struct
+ type quantifier = Quantifier of expr
+
+ let expr_of_quantifier e = match e with Quantifier(x) -> x
+
+ let quantifier_of_expr e =
+ match e with Expr.Expr(a) ->
+ let q = (Z3enums.ast_kind_of_int (Z3native.get_ast_kind (z3obj_gnc a) (z3obj_gno a))) in
+ if (q != Z3enums.QUANTIFIER_AST) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ Quantifier(e)
+
+ let gc ( x : quantifier ) = match (x) with Quantifier(e) -> (Expr.gc e)
+ let gnc ( x : quantifier ) = match (x) with Quantifier(e) -> (Expr.gnc e)
+ let gno ( x : quantifier ) = match (x) with Quantifier(e) -> (Expr.gno e)
+
+ module Pattern =
+ struct
+ type pattern = Pattern of ast
+
+ let ast_of_pattern e = match e with Pattern(x) -> x
+
+ let pattern_of_ast a =
+ (* CMW: Unchecked ok? *)
+ Pattern(a)
+
+ let gc ( x : pattern ) = match (x) with Pattern(a) -> (z3obj_gc a)
+ let gnc ( x : pattern ) = match (x) with Pattern(a) -> (z3obj_gnc a)
+ let gno ( x : pattern ) = match (x) with Pattern(a) -> (z3obj_gno a)
+
+ let get_num_terms ( x : pattern ) =
+ Z3native.get_pattern_num_terms (gnc x) (gno x)
+
+ let get_terms ( x : pattern ) =
+ let n = (get_num_terms x) in
+ let f i = (expr_of_ptr (gc x) (Z3native.get_pattern (gnc x) (gno x) i)) in
+ mk_list f n
+
+ let to_string ( x : pattern ) = Z3native.pattern_to_string (gnc x) (gno x)
+ end
+
+ let get_index ( x : expr ) =
+ if not (AST.is_var (match x with Expr.Expr(a) -> a)) then
+ raise (Z3native.Exception "Term is not a bound variable.")
+ else
+ Z3native.get_index_value (Expr.gnc x) (Expr.gno x)
+
+ let is_universal ( x : quantifier ) =
+ Z3native.is_quantifier_forall (gnc x) (gno x)
+
+ let is_existential ( x : quantifier ) = not (is_universal x)
+
+ let get_weight ( x : quantifier ) = Z3native.get_quantifier_weight (gnc x) (gno x)
+
+ let get_num_patterns ( x : quantifier ) = Z3native.get_quantifier_num_patterns (gnc x) (gno x)
+
+ let get_patterns ( x : quantifier ) =
+ let n = (get_num_patterns x) in
+ let f i = Pattern.Pattern (z3_native_object_of_ast_ptr (gc x) (Z3native.get_quantifier_pattern_ast (gnc x) (gno x) i)) in
+ mk_list f n
+
+ let get_num_no_patterns ( x : quantifier ) = Z3native.get_quantifier_num_no_patterns (gnc x) (gno x)
+
+ let get_no_patterns ( x : quantifier ) =
+ let n = (get_num_patterns x) in
+ let f i = Pattern.Pattern (z3_native_object_of_ast_ptr (gc x) (Z3native.get_quantifier_no_pattern_ast (gnc x) (gno x) i)) in
+ mk_list f n
+
+ let get_num_bound ( x : quantifier ) = Z3native.get_quantifier_num_bound (gnc x) (gno x)
+
+ let get_bound_variable_names ( x : quantifier ) =
+ let n = (get_num_bound x) in
+ let f i = (Symbol.create (gc x) (Z3native.get_quantifier_bound_name (gnc x) (gno x) i)) in
+ mk_list f n
+
+ let get_bound_variable_sorts ( x : quantifier ) =
+ let n = (get_num_bound x) in
+ let f i = (sort_of_ptr (gc x) (Z3native.get_quantifier_bound_sort (gnc x) (gno x) i)) in
+ mk_list f n
+
+ let get_body ( x : quantifier ) =
+ Boolean.bool_expr_of_ptr (gc x) (Z3native.get_quantifier_body (gnc x) (gno x))
+
+ let mk_bound ( ctx : context ) ( index : int ) ( ty : sort ) =
+ expr_of_ptr ctx (Z3native.mk_bound (context_gno ctx) index (Sort.gno ty))
+
+ let mk_pattern ( ctx : context ) ( terms : expr list ) =
+ if (List.length terms) == 0 then
+ raise (Z3native.Exception "Cannot create a pattern from zero terms")
+ else
+ Pattern.Pattern(z3_native_object_of_ast_ptr ctx (Z3native.mk_pattern (context_gno ctx) (List.length terms) (expr_lton terms)))
+
+ let mk_forall ( ctx : context ) ( sorts : sort list ) ( names : Symbol.symbol list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
+ if (List.length sorts) != (List.length names) then
+ raise (Z3native.Exception "Number of sorts does not match number of names")
+ else if ((List.length nopatterns) == 0 && quantifier_id == None && skolem_id == None) then
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier (context_gno ctx) true
+ (match weight with | None -> 1 | Some(x) -> x)
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (List.length sorts) (sort_lton sorts)
+ (Symbol.symbol_lton names)
+ (Expr.gno body)))
+ else
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_ex (context_gno ctx) true
+ (match weight with | None -> 1 | Some(x) -> x)
+ (match quantifier_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (match skolem_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (List.length nopatterns) (expr_lton nopatterns)
+ (List.length sorts) (sort_lton sorts)
+ (Symbol.symbol_lton names)
+ (Expr.gno body)))
+
+ let mk_forall_const ( ctx : context ) ( bound_constants : expr list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
+ if ((List.length nopatterns) == 0 && quantifier_id == None && skolem_id == None) then
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_const (context_gno ctx) true
+ (match weight with | None -> 1 | Some(x) -> x)
+ (List.length bound_constants) (expr_lton bound_constants)
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (Expr.gno body)))
+ else
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_const_ex (context_gno ctx) true
+ (match weight with | None -> 1 | Some(x) -> x)
+ (match quantifier_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (match skolem_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (List.length bound_constants) (expr_lton bound_constants)
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (List.length nopatterns) (expr_lton nopatterns)
+ (Expr.gno body)))
+
+ let mk_exists ( ctx : context ) ( sorts : sort list ) ( names : Symbol.symbol list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
+ if (List.length sorts) != (List.length names) then
+ raise (Z3native.Exception "Number of sorts does not match number of names")
+ else if ((List.length nopatterns) == 0 && quantifier_id == None && skolem_id == None) then
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier (context_gno ctx) false
+ (match weight with | None -> 1 | Some(x) -> x)
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (List.length sorts) (sort_lton sorts)
+ (Symbol.symbol_lton names)
+ (Expr.gno body)))
+ else
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_ex (context_gno ctx) false
+ (match weight with | None -> 1 | Some(x) -> x)
+ (match quantifier_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (match skolem_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (List.length nopatterns) (expr_lton nopatterns)
+ (List.length sorts) (sort_lton sorts)
+ (Symbol.symbol_lton names)
+ (Expr.gno body)))
+
+ let mk_exists_const ( ctx : context ) ( bound_constants : expr list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
+ if ((List.length nopatterns) == 0 && quantifier_id == None && skolem_id == None) then
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_const (context_gno ctx) false
+ (match weight with | None -> 1 | Some(x) -> x)
+ (List.length bound_constants) (expr_lton bound_constants)
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (Expr.gno body)))
+ else
+ Quantifier(expr_of_ptr ctx (Z3native.mk_quantifier_const_ex (context_gno ctx) false
+ (match weight with | None -> 1 | Some(x) -> x)
+ (match quantifier_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (match skolem_id with | None -> null | Some(x) -> (Symbol.gno x))
+ (List.length bound_constants) (expr_lton bound_constants)
+ (List.length patterns) (let f x = (AST.ptr_of_ast (Pattern.ast_of_pattern x)) in (Array.of_list (List.map f patterns)))
+ (List.length nopatterns) (expr_lton nopatterns)
+ (Expr.gno body)))
+
+ let mk_quantifier ( ctx : context ) ( universal : bool ) ( sorts : sort list ) ( names : Symbol.symbol list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
+ if (universal) then
+ (mk_forall ctx sorts names body weight patterns nopatterns quantifier_id skolem_id)
+ else
+ (mk_exists ctx sorts names body weight patterns nopatterns quantifier_id skolem_id)
+
+ let mk_quantifier ( ctx : context ) ( universal : bool ) ( bound_constants : expr list ) ( body : expr ) ( weight : int option ) ( patterns : Pattern.pattern list ) ( nopatterns : expr list ) ( quantifier_id : Symbol.symbol option ) ( skolem_id : Symbol.symbol option ) =
+ if (universal) then
+ mk_forall_const ctx bound_constants body weight patterns nopatterns quantifier_id skolem_id
+ else
+ mk_exists_const ctx bound_constants body weight patterns nopatterns quantifier_id skolem_id
+end
+
+
+module Array_ =
+struct
+ type array_sort = ArraySort of sort
+ type array_expr = ArrayExpr of expr
+
+ let array_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let e = (expr_of_ptr ctx no) in
+ ArrayExpr(e)
+
+ let array_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let s = (sort_of_ptr ctx no) in
+ ArraySort(s)
+
+ let sort_of_array_sort s = match s with ArraySort(x) -> x
+
+ let array_sort_of_sort s = match s with Sort(a) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.ARRAY_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ ArraySort(s)
+
+ let array_expr_of_expr e =
+ match e with Expr(a) ->
+ let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
+ if (q != Z3enums.ARRAY_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ ArrayExpr(e)
+
+ let expr_of_array_expr e = match e with ArrayExpr(x) -> x
+
+ let sgc ( x : array_sort ) = match (x) with ArraySort(Sort(s)) -> (z3obj_gc s)
+ let sgnc ( x : array_sort ) = match (x) with ArraySort(Sort(s)) -> (z3obj_gnc s)
+ let sgno ( x : array_sort ) = match (x) with ArraySort(Sort(s)) -> (z3obj_gno s)
+
+ let egc ( x : array_expr ) = match (x) with ArrayExpr(Expr(e)) -> (z3obj_gc e)
+ let egnc ( x : array_expr ) = match (x) with ArrayExpr(Expr(e)) -> (z3obj_gnc e)
+ let egno ( x : array_expr ) = match (x) with ArrayExpr(Expr(e)) -> (z3obj_gno e)
+
+ let mk_sort ( ctx : context ) ( domain : sort ) ( range : sort ) =
+ array_sort_of_ptr ctx (Z3native.mk_array_sort (context_gno ctx) (Sort.gno domain) (Sort.gno range))
+
+ let is_store ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_STORE)
+ let is_select ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SELECT)
+ let is_constant_array ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_CONST_ARRAY)
+ let is_default_array ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ARRAY_DEFAULT)
+ let is_array_map ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ARRAY_MAP)
+ let is_as_array ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_AS_ARRAY)
+ let is_array ( x : expr ) =
+ (Z3native.is_app (Expr.gnc x) (Expr.gno x)) &&
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == ARRAY_SORT)
+
+ let get_domain ( x : array_sort ) = Sort.sort_of_ptr (sgc x) (Z3native.get_array_sort_domain (sgnc x) (sgno x))
+ let get_range ( x : array_sort ) = Sort.sort_of_ptr (sgc x) (Z3native.get_array_sort_range (sgnc x) (sgno x))
+
+ let mk_const ( ctx : context ) ( name : Symbol.symbol ) ( domain : sort ) ( range : sort ) =
+ ArrayExpr(Expr.mk_const ctx name (match (mk_sort ctx domain range) with ArraySort(s) -> s))
+
+ let mk_const_s ( ctx : context ) ( name : string ) ( domain : sort ) ( range : sort ) =
+ mk_const ctx (Symbol.mk_string ctx name) domain range
+
+ let mk_select ( ctx : context ) ( a : array_expr ) ( i : expr ) =
+ array_expr_of_ptr ctx (Z3native.mk_select (context_gno ctx) (egno a) (Expr.gno i))
+
+ let mk_store ( ctx : context ) ( a : array_expr ) ( i : expr ) ( v : expr ) =
+ array_expr_of_ptr ctx (Z3native.mk_store (context_gno ctx) (egno a) (Expr.gno i) (Expr.gno v))
+
+ let mk_const_array ( ctx : context ) ( domain : sort ) ( v : expr ) =
+ array_expr_of_ptr ctx (Z3native.mk_const_array (context_gno ctx) (Sort.gno domain) (Expr.gno v))
+
+ let mk_map ( ctx : context ) ( f : func_decl ) ( args : array_expr list ) =
+ let m x = (Expr.gno (expr_of_array_expr x)) in
+ array_expr_of_ptr ctx (Z3native.mk_map (context_gno ctx) (FuncDecl.gno f) (List.length args) (Array.of_list (List.map m args)))
+
+ let mk_term_array ( ctx : context ) ( arg : array_expr ) =
+ array_expr_of_ptr ctx (Z3native.mk_array_default (context_gno ctx) (egno arg))
+end
+
+
+module Set =
+struct
+ type set_sort = SetSort of sort
+
+ let set_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let s = (sort_of_ptr ctx no) in
+ SetSort(s)
+
+ let sort_of_set_sort s = match s with SetSort(x) -> x
+
+ let mk_sort ( ctx : context ) ( ty : sort ) =
+ set_sort_of_ptr ctx (Z3native.mk_set_sort (context_gno ctx) (Sort.gno ty))
+
+ let is_union ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SET_UNION)
+ let is_intersect ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SET_INTERSECT)
+ let is_difference ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SET_DIFFERENCE)
+ let is_complement ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SET_COMPLEMENT)
+ let is_subset ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SET_SUBSET)
+
+
+ let mk_empty ( ctx : context ) ( domain : sort ) =
+ (expr_of_ptr ctx (Z3native.mk_empty_set (context_gno ctx) (Sort.gno domain)))
+
+ let mk_full ( ctx : context ) ( domain : sort ) =
+ expr_of_ptr ctx (Z3native.mk_full_set (context_gno ctx) (Sort.gno domain))
+
+ let mk_set_add ( ctx : context ) ( set : expr ) ( element : expr ) =
+ expr_of_ptr ctx (Z3native.mk_set_add (context_gno ctx) (Expr.gno set) (Expr.gno element))
+
+ let mk_del ( ctx : context ) ( set : expr ) ( element : expr ) =
+ expr_of_ptr ctx (Z3native.mk_set_del (context_gno ctx) (Expr.gno set) (Expr.gno element))
+
+ let mk_union ( ctx : context ) ( args : expr list ) =
+ expr_of_ptr ctx (Z3native.mk_set_union (context_gno ctx) (List.length args) (expr_lton args))
+
+ let mk_intersection ( ctx : context ) ( args : expr list ) =
+ expr_of_ptr ctx (Z3native.mk_set_intersect (context_gno ctx) (List.length args) (expr_lton args))
+
+ let mk_difference ( ctx : context ) ( arg1 : expr ) ( arg2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_set_difference (context_gno ctx) (Expr.gno arg1) (Expr.gno arg2))
+
+ let mk_complement ( ctx : context ) ( arg : expr ) =
+ expr_of_ptr ctx (Z3native.mk_set_complement (context_gno ctx) (Expr.gno arg))
+
+ let mk_membership ( ctx : context ) ( elem : expr ) ( set : expr ) =
+ expr_of_ptr ctx (Z3native.mk_set_member (context_gno ctx) (Expr.gno elem) (Expr.gno set))
+
+ let mk_subset ( ctx : context ) ( arg1 : expr ) ( arg2 : expr ) =
+ expr_of_ptr ctx (Z3native.mk_set_subset (context_gno ctx) (Expr.gno arg1) (Expr.gno arg2))
+
+end
+
+
+module FiniteDomain =
+struct
+ type finite_domain_sort = FiniteDomainSort of sort
+
+ let sort_of_finite_domain_sort s = match s with FiniteDomainSort(x) -> x
+
+ let finite_domain_sort_of_sort s = match s with Sort(a) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.FINITE_DOMAIN_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ FiniteDomainSort(s)
+
+ let gc ( x : finite_domain_sort ) = match (x) with FiniteDomainSort(Sort(s)) -> (z3obj_gc s)
+ let gnc ( x : finite_domain_sort ) = match (x) with FiniteDomainSort(Sort(s)) -> (z3obj_gnc s)
+ let gno ( x : finite_domain_sort ) = match (x) with FiniteDomainSort(Sort(s))-> (z3obj_gno s)
+
+ let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( size : int ) =
+ let s = (sort_of_ptr ctx (Z3native.mk_finite_domain_sort (context_gno ctx) (Symbol.gno name) size)) in
+ FiniteDomainSort(s)
+
+ let mk_sort_s ( ctx : context ) ( name : string ) ( size : int ) =
+ mk_sort ctx (Symbol.mk_string ctx name) size
+
+
+ let is_finite_domain ( x : expr ) =
+ let nc = (Expr.gnc x) in
+ (Z3native.is_app (Expr.gnc x) (Expr.gno x)) &&
+ (sort_kind_of_int (Z3native.get_sort_kind nc (Z3native.get_sort nc (Expr.gno x))) == FINITE_DOMAIN_SORT)
+
+ let is_lt ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_FD_LT)
+
+ let get_size ( x : finite_domain_sort ) =
+ let (r, v) = (Z3native.get_finite_domain_sort_size (gnc x) (gno x)) in
+ if r then v
+ else raise (Z3native.Exception "Conversion failed.")
+end
+
+
+module Relation =
+struct
+ type relation_sort = RelationSort of sort
+
+ let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let s = (sort_of_ptr ctx no) in
+ RelationSort(s)
+
+ let sort_of_relation_sort s = match s with RelationSort(x) -> x
+
+ let relation_sort_of_sort s = match s with Sort(a) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.RELATION_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ RelationSort(s)
+
+ let gc ( x : relation_sort ) = match (x) with RelationSort(Sort(s)) -> (z3obj_gc s)
+ let gnc ( x : relation_sort ) = match (x) with RelationSort(Sort(s)) -> (z3obj_gnc s)
+ let gno ( x : relation_sort ) = match (x) with RelationSort(Sort(s))-> (z3obj_gno s)
+
+
+ let is_relation ( x : expr ) =
+ let nc = (Expr.gnc x) in
+ ((Z3native.is_app (Expr.gnc x) (Expr.gno x)) &&
+ (sort_kind_of_int (Z3native.get_sort_kind nc (Z3native.get_sort nc (Expr.gno x))) == RELATION_SORT))
+
+ let is_store ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_STORE)
+ let is_empty ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_EMPTY)
+ let is_is_empty ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_IS_EMPTY)
+ let is_join ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_JOIN)
+ let is_union ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_UNION)
+ let is_widen ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_WIDEN)
+ let is_project ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_PROJECT)
+ let is_filter ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_FILTER)
+ let is_negation_filter ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_NEGATION_FILTER)
+ let is_rename ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_RENAME)
+ let is_complement ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_COMPLEMENT)
+ let is_select ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_SELECT)
+ let is_clone ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_RA_CLONE)
+
+ let get_arity ( x : relation_sort ) = Z3native.get_relation_arity (gnc x) (gno x)
+
+ let get_column_sorts ( x : relation_sort ) =
+ let n = get_arity x in
+ let f i = (sort_of_ptr (gc x) (Z3native.get_relation_column (gnc x) (gno x) i)) in
+ mk_list f n
+
+end
+
+
+module Datatype =
+struct
+ type datatype_sort = DatatypeSort of sort
+ type datatype_expr = DatatypeExpr of expr
+
+ let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let s = (sort_of_ptr ctx no) in
+ DatatypeSort(s)
+
+ let sort_of_datatype_sort s = match s with DatatypeSort(x) -> x
+
+ let datatype_sort_of_sort s = match s with Sort(a) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.DATATYPE_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ DatatypeSort(s)
+
+ let datatype_expr_of_expr e =
+ match e with Expr(a) ->
+ let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
+ if (q != Z3enums.DATATYPE_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ DatatypeExpr(e)
+
+ let expr_of_datatype_expr e = match e with DatatypeExpr(x) -> x
+
+ let sgc ( x : datatype_sort ) = match (x) with DatatypeSort(Sort(s)) -> (z3obj_gc s)
+ let sgnc ( x : datatype_sort ) = match (x) with DatatypeSort(Sort(s)) -> (z3obj_gnc s)
+ let sgno ( x : datatype_sort ) = match (x) with DatatypeSort(Sort(s))-> (z3obj_gno s)
+
+ module Constructor =
+ struct
+ type constructor = z3_native_object
+
+ let _sizes = Hashtbl.create 0
+
+ let create ( ctx : context ) ( name : Symbol.symbol ) ( recognizer : Symbol.symbol ) ( field_names : Symbol.symbol list ) ( sorts : sort list ) ( sort_refs : int list ) =
+ let n = (List.length field_names) in
+ if n != (List.length sorts) then
+ raise (Z3native.Exception "Number of field names does not match number of sorts")
+ else
+ if n != (List.length sort_refs) then
+ raise (Z3native.Exception "Number of field names does not match number of sort refs")
+ else
+ let ptr = (Z3native.mk_constructor (context_gno ctx) (Symbol.gno name)
+ (Symbol.gno recognizer)
+ n
+ (Symbol.symbol_lton field_names)
+ (sort_lton sorts)
+ (Array.of_list sort_refs)) in
+ let no : constructor = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = z3obj_nil_ref ;
+ dec_ref = z3obj_nil_ref} in
+ (z3obj_sno no ctx ptr) ;
+ (z3obj_create no) ;
+ let f = fun o -> Z3native.del_constructor (z3obj_gnc o) (z3obj_gno o) in
+ Gc.finalise f no ;
+ Hashtbl.add _sizes no n ;
+ no
+
+ let get_num_fields ( x : constructor ) = Hashtbl.find _sizes x
+
+ let get_constructor_decl ( x : constructor ) =
+ let (a, _, _) = (Z3native.query_constructor (z3obj_gnc x) (z3obj_gno x) (get_num_fields x)) in
+ func_decl_of_ptr (z3obj_gc x) a
+
+ let get_tester_decl ( x : constructor ) =
+ let (_, b, _) = (Z3native.query_constructor (z3obj_gnc x) (z3obj_gno x) (get_num_fields x)) in
+ func_decl_of_ptr (z3obj_gc x) b
+
+ let get_accessor_decls ( x : constructor ) =
+ let (_, _, c) = (Z3native.query_constructor (z3obj_gnc x) (z3obj_gno x) (get_num_fields x)) in
+ let f i = func_decl_of_ptr (z3obj_gc x) (Array.get c i) in
+ mk_list f (Array.length c)
+
+ end
+
+ module ConstructorList =
+ struct
+ type constructor_list = z3_native_object
+
+ let create ( ctx : context ) ( c : Constructor.constructor list ) =
+ let res : constructor_list = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = z3obj_nil_ref ;
+ dec_ref = z3obj_nil_ref} in
+ let f x =(z3obj_gno x) in
+ (z3obj_sno res ctx (Z3native.mk_constructor_list (context_gno ctx) (List.length c) (Array.of_list (List.map f c)))) ;
+ (z3obj_create res) ;
+ let f = fun o -> Z3native.del_constructor_list (z3obj_gnc o) (z3obj_gno o) in
+ Gc.finalise f res;
+ res
+ end
+
+ let mk_constructor ( ctx : context ) ( name : Symbol.symbol ) ( recognizer : Symbol.symbol ) ( field_names : Symbol.symbol list ) ( sorts : sort list ) ( sort_refs : int list ) =
+ Constructor.create ctx name recognizer field_names sorts sort_refs
+
+
+ let mk_constructor_s ( ctx : context ) ( name : string ) ( recognizer : Symbol.symbol ) ( field_names : Symbol.symbol list ) ( sorts : sort list ) ( sort_refs : int list ) =
+ mk_constructor ctx (Symbol.mk_string ctx name) recognizer field_names sorts sort_refs
+
+ let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( constructors : Constructor.constructor list ) =
+ let f x = (z3obj_gno x) in
+ let (x,_) = (Z3native.mk_datatype (context_gno ctx) (Symbol.gno name) (List.length constructors) (Array.of_list (List.map f constructors))) in
+ sort_of_ptr ctx x
+
+ let mk_sort_s ( ctx : context ) ( name : string ) ( constructors : Constructor.constructor list ) =
+ mk_sort ctx (Symbol.mk_string ctx name) constructors
+
+ let mk_sorts ( ctx : context ) ( names : Symbol.symbol list ) ( c : Constructor.constructor list list ) =
+ let n = (List.length names) in
+ let f e = (AST.ptr_of_ast (ConstructorList.create ctx e)) in
+ let cla = (Array.of_list (List.map f c)) in
+ let (r, a) = (Z3native.mk_datatypes (context_gno ctx) n (Symbol.symbol_lton names) cla) in
+ let g i = (sort_of_ptr ctx (Array.get r i)) in
+ mk_list g (Array.length r)
+
+ let mk_sorts_s ( ctx : context ) ( names : string list ) ( c : Constructor.constructor list list ) =
+ mk_sorts ctx
+ (
+ let f e = (Symbol.mk_string ctx e) in
+ List.map f names
+ )
+ c
+
+ let get_num_constructors ( x : datatype_sort ) = Z3native.get_datatype_sort_num_constructors (sgnc x) (sgno x)
+
+ let get_constructors ( x : datatype_sort ) =
+ let n = (get_num_constructors x) in
+ let f i = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) i) in
+ mk_list f n
+
+ let get_recognizers ( x : datatype_sort ) =
+ let n = (get_num_constructors x) in
+ let f i = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) i) in
+ mk_list f n
+
+ let get_accessors ( x : datatype_sort ) =
+ let n = (get_num_constructors x) in
+ let f i = (
+ let fd = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) i) in
+ let ds = Z3native.get_domain_size (FuncDecl.gnc fd) (FuncDecl.gno fd) in
+ let g j = func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor_accessor (sgnc x) (sgno x) i j) in
+ mk_list g ds
+ ) in
+ mk_list f n
+end
+
+
+module Enumeration =
+struct
+ type enum_sort = EnumSort of sort
+
+ let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) ( cdecls : Z3native.z3_func_decl list ) ( tdecls : Z3native.z3_func_decl list ) =
+ let s = (sort_of_ptr ctx no) in
+ let res = EnumSort(s) in
+ res
+
+ let sort_of_enum_sort s = match s with EnumSort(x) -> x
+
+ let sgc ( x : enum_sort ) = match (x) with EnumSort(Sort(s)) -> (z3obj_gc s)
+ let sgnc ( x : enum_sort ) = match (x) with EnumSort(Sort(s)) -> (z3obj_gnc s)
+ let sgno ( x : enum_sort ) = match (x) with EnumSort(Sort(s))-> (z3obj_gno s)
+
+ let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( enum_names : Symbol.symbol list ) =
+ let (a, b, c) = (Z3native.mk_enumeration_sort (context_gno ctx) (Symbol.gno name) (List.length enum_names) (Symbol.symbol_lton enum_names)) in
+ sort_of_ptr ctx a (list_of_array b) (list_of_array c)
+
+ let mk_sort_s ( ctx : context ) ( name : string ) ( enum_names : string list ) =
+ mk_sort ctx (Symbol.mk_string ctx name) (Symbol.mk_strings ctx enum_names)
+
+ let get_const_decls ( x : enum_sort ) =
+ let n = Z3native.get_datatype_sort_num_constructors (sgnc x) (sgno x) in
+ let f i = (func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) i)) in
+ mk_list f n
+
+ let get_tester_decls ( x : enum_sort ) =
+ let n = Z3native.get_datatype_sort_num_constructors (sgnc x) (sgno x) in
+ let f i = (func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) i)) in
+ mk_list f n
+
+end
+
+
+module List_ =
+struct
+ type list_sort = ListSort of sort
+
+ let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) ( nildecl : Z3native.ptr ) ( is_nildecl : Z3native.ptr ) ( consdecl : Z3native.ptr ) ( is_consdecl : Z3native.ptr ) ( headdecl : Z3native.ptr ) ( taildecl : Z3native.ptr ) =
+ let s = (sort_of_ptr ctx no) in
+ let res = ListSort(s) in
+ res
+
+ let sort_of_list_sort s = match s with ListSort(x) -> x
+
+ let sgc ( x : list_sort ) = match (x) with ListSort(Sort(s)) -> (z3obj_gc s)
+ let sgnc ( x : list_sort ) = match (x) with ListSort(Sort(s)) -> (z3obj_gnc s)
+ let sgno ( x : list_sort ) = match (x) with ListSort(Sort(s))-> (z3obj_gno s)
+
+ let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( elem_sort : sort ) =
+ let (r, a, b, c, d, e, f) = (Z3native.mk_list_sort (context_gno ctx) (Symbol.gno name) (Sort.gno elem_sort)) in
+ sort_of_ptr ctx r a b c d e f
+
+ let mk_list_s ( ctx : context ) (name : string) elem_sort =
+ mk_sort ctx (Symbol.mk_string ctx name) elem_sort
+
+ let get_nil_decl ( x : list_sort ) =
+ func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) 0)
+
+ let get_is_nil_decl ( x : list_sort ) =
+ func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) 0)
+
+ let get_cons_decl ( x : list_sort ) =
+ func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor (sgnc x) (sgno x) 1)
+
+ let get_is_cons_decl ( x : list_sort ) =
+ func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_recognizer (sgnc x) (sgno x) 1)
+
+ let get_head_decl ( x : list_sort ) =
+ func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor_accessor (sgnc x) (sgno x) 1 0)
+
+ let get_tail_decl ( x : list_sort ) =
+ func_decl_of_ptr (sgc x) (Z3native.get_datatype_sort_constructor_accessor (sgnc x) (sgno x) 1 1)
+
+ let nil ( x : list_sort ) = expr_of_func_app (sgc x) (get_nil_decl x) []
+end
+
+
+module Tuple =
+struct
+ type tuple_sort = TupleSort of sort
+
+ let sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ let s = (sort_of_ptr ctx no) in
+ TupleSort(s)
+
+ let sort_of_tuple_sort s = match s with TupleSort(x) -> x
+
+ let sgc ( x : tuple_sort ) = match (x) with TupleSort(Sort(s)) -> (z3obj_gc s)
+ let sgnc ( x : tuple_sort ) = match (x) with TupleSort(Sort(s)) -> (z3obj_gnc s)
+ let sgno ( x : tuple_sort ) = match (x) with TupleSort(Sort(s))-> (z3obj_gno s)
+
+ let mk_sort ( ctx : context ) ( name : Symbol.symbol ) ( field_names : Symbol.symbol list ) ( field_sorts : sort list ) =
+ let (r, _, _) = (Z3native.mk_tuple_sort (context_gno ctx) (Symbol.gno name) (List.length field_names) (Symbol.symbol_lton field_names) (sort_lton field_sorts)) in
+ sort_of_ptr ctx r
+
+ let get_mk_decl ( x : tuple_sort ) =
+ func_decl_of_ptr (sgc x) (Z3native.get_tuple_sort_mk_decl (sgnc x) (sgno x))
+
+ let get_num_fields ( x : tuple_sort ) = Z3native.get_tuple_sort_num_fields (sgnc x) (sgno x)
+
+ let get_field_decls ( x : tuple_sort ) =
+ let n = get_num_fields x in
+ let f i = func_decl_of_ptr (sgc x) (Z3native.get_tuple_sort_field_decl (sgnc x) (sgno x) i) in
+ mk_list f n
+end
+
+
+module rec Arithmetic :
+sig
+ type arith_sort = ArithSort of Sort.sort
+ type arith_expr = ArithExpr of Expr.expr
+
+ val sort_of_arith_sort : arith_sort -> Sort.sort
+ val arith_sort_of_sort : Sort.sort -> arith_sort
+ val expr_of_arith_expr : arith_expr -> Expr.expr
+ val arith_expr_of_expr : Expr.expr -> arith_expr
+
+ module rec Integer :
+ sig
+ type int_sort = IntSort of arith_sort
+ type int_expr = IntExpr of arith_expr
+ type int_num = IntNum of int_expr
+
+ val int_expr_of_ptr : context -> Z3native.ptr -> int_expr
+ val int_num_of_ptr : context -> Z3native.ptr -> int_num
+
+ val arith_sort_of_int_sort : Integer.int_sort -> arith_sort
+ val int_sort_of_arith_sort : arith_sort -> int_sort
+ val arith_expr_of_int_expr : int_expr -> arith_expr
+ val int_expr_of_int_num : int_num -> int_expr
+ val int_expr_of_arith_expr : arith_expr -> int_expr
+ val int_num_of_int_expr : int_expr -> int_num
+
+ val mk_sort : context -> int_sort
+ val get_int : int_num -> int
+ val to_string : int_num -> string
+ val mk_int_const : context -> Symbol.symbol -> int_expr
+ val mk_int_const_s : context -> string -> int_expr
+ val mk_mod : context -> int_expr -> int_expr -> int_expr
+ val mk_rem : context -> int_expr -> int_expr -> int_expr
+ val mk_int_numeral_s : context -> string -> int_num
+ val mk_int_numeral_i : context -> int -> int_num
+ val mk_int2real : context -> int_expr -> Real.real_expr
+ val mk_int2bv : context -> int -> int_expr -> BitVector.bitvec_expr
+ end
+ and Real :
+ sig
+ type real_sort = RealSort of arith_sort
+ type real_expr = RealExpr of arith_expr
+ type rat_num = RatNum of real_expr
+
+ val real_expr_of_ptr : context -> Z3native.ptr -> real_expr
+ val rat_num_of_ptr : context -> Z3native.ptr -> rat_num
+
+ val arith_sort_of_real_sort : Arithmetic.Real.real_sort -> Arithmetic.arith_sort
+ val real_sort_of_arith_sort : Arithmetic.arith_sort -> Arithmetic.Real.real_sort
+ val arith_expr_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.arith_expr
+ val real_expr_of_rat_num : Arithmetic.Real.rat_num -> Arithmetic.Real.real_expr
+ val real_expr_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.Real.real_expr
+ val rat_num_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.Real.rat_num
+
+ val mk_sort : context -> real_sort
+ val get_numerator : rat_num -> Integer.int_num
+ val get_denominator : rat_num -> Integer.int_num
+ val to_decimal_string : rat_num -> int -> string
+ val to_string : rat_num -> string
+ val mk_real_const : context -> Symbol.symbol -> real_expr
+ val mk_real_const_s : context -> string -> real_expr
+ val mk_numeral_nd : context -> int -> int -> rat_num
+ val mk_numeral_s : context -> string -> rat_num
+ val mk_numeral_i : context -> int -> rat_num
+ val mk_is_integer : context -> real_expr -> Boolean.bool_expr
+ val mk_real2int : context -> real_expr -> Integer.int_expr
+ end
+ and AlgebraicNumber :
+ sig
+ type algebraic_num = AlgebraicNum of arith_expr
+
+ val arith_expr_of_algebraic_num : algebraic_num -> arith_expr
+ val algebraic_num_of_arith_expr : arith_expr -> algebraic_num
+
+ val to_upper : algebraic_num -> int -> Real.rat_num
+ val to_lower : algebraic_num -> int -> Real.rat_num
+ val to_decimal_string : algebraic_num -> int -> string
+ val to_string : algebraic_num -> string
+ end
+
+ val is_int : Expr.expr -> bool
+ val is_arithmetic_numeral : Expr.expr -> bool
+ val is_le : Expr.expr -> bool
+ val is_ge : Expr.expr -> bool
+ val is_lt : Expr.expr -> bool
+ val is_gt : Expr.expr -> bool
+ val is_add : Expr.expr -> bool
+ val is_sub : Expr.expr -> bool
+ val is_uminus : Expr.expr -> bool
+ val is_mul : Expr.expr -> bool
+ val is_div : Expr.expr -> bool
+ val is_idiv : Expr.expr -> bool
+ val is_remainder : Expr.expr -> bool
+ val is_modulus : Expr.expr -> bool
+ val is_inttoreal : Expr.expr -> bool
+ val is_real_to_int : Expr.expr -> bool
+ val is_real_is_int : Expr.expr -> bool
+ val is_real : Expr.expr -> bool
+ val is_int_numeral : Expr.expr -> bool
+ val is_rat_num : Expr.expr -> bool
+ val is_algebraic_number : Expr.expr -> bool
+ val mk_add : context -> arith_expr list -> arith_expr
+ val mk_mul : context -> arith_expr list -> arith_expr
+ val mk_sub : context -> arith_expr list -> arith_expr
+ val mk_unary_minus : context -> arith_expr -> arith_expr
+ val mk_div : context -> arith_expr -> arith_expr -> arith_expr
+ val mk_power : context -> arith_expr -> arith_expr -> arith_expr
+ val mk_lt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+ val mk_le : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+ val mk_gt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+ val mk_ge : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+end = struct
+ type arith_sort = ArithSort of sort
+ type arith_expr = ArithExpr of expr
+
+ let arith_expr_of_expr e =
+ match e with Expr(a) ->
+ let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
+ if (q != Z3enums.INT_SORT && q != Z3enums.REAL_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ ArithExpr(e)
+
+ let arith_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ arith_expr_of_expr (expr_of_ptr ctx no)
+
+ let sort_of_arith_sort s = match s with ArithSort(x) -> x
+ let expr_of_arith_expr e = match e with ArithExpr(x) -> x
+
+ let arith_sort_of_sort s = match s with Sort(a) ->
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) in
+ if (q != Z3enums.INT_SORT && q != Z3enums.REAL_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ ArithSort(s)
+
+ let arith_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ arith_sort_of_sort (sort_of_ptr ctx no)
+
+ let sgc ( x : arith_sort ) = match (x) with ArithSort(Sort(s)) -> (z3obj_gc s)
+ let sgnc ( x : arith_sort ) = match (x) with ArithSort(Sort(s)) -> (z3obj_gnc s)
+ let sgno ( x : arith_sort ) = match (x) with ArithSort(Sort(s)) -> (z3obj_gno s)
+ let egc ( x : arith_expr ) = match (x) with ArithExpr(e) -> (Expr.gc e)
+ let egnc ( x : arith_expr ) = match (x) with ArithExpr(e) -> (Expr.gnc e)
+ let egno ( x : arith_expr ) = match (x) with ArithExpr(e) -> (Expr.gno e)
+
+ module rec Integer :
+ sig
+ type int_sort = IntSort of arith_sort
+ type int_expr = IntExpr of arith_expr
+ type int_num = IntNum of int_expr
+
+ val int_expr_of_ptr : context -> Z3native.ptr -> int_expr
+ val int_num_of_ptr : context -> Z3native.ptr -> int_num
+
+ val arith_sort_of_int_sort : Integer.int_sort -> arith_sort
+ val int_sort_of_arith_sort : arith_sort -> int_sort
+ val arith_expr_of_int_expr : int_expr -> arith_expr
+ val int_expr_of_int_num : int_num -> int_expr
+ val int_expr_of_arith_expr : arith_expr -> int_expr
+ val int_num_of_int_expr : int_expr -> int_num
+
+ val mk_sort : context -> int_sort
+ val get_int : int_num -> int
+ val to_string : int_num -> string
+ val mk_int_const : context -> Symbol.symbol -> int_expr
+ val mk_int_const_s : context -> string -> int_expr
+ val mk_mod : context -> int_expr -> int_expr -> int_expr
+ val mk_rem : context -> int_expr -> int_expr -> int_expr
+ val mk_int_numeral_s : context -> string -> int_num
+ val mk_int_numeral_i : context -> int -> int_num
+ val mk_int2real : context -> int_expr -> Real.real_expr
+ val mk_int2bv : context -> int -> int_expr -> BitVector.bitvec_expr
+ end = struct
+ type int_sort = IntSort of arith_sort
+ type int_expr = IntExpr of arith_expr
+ type int_num = IntNum of int_expr
+
+ let int_expr_of_arith_expr e =
+ match e with ArithExpr(Expr(a)) ->
+ let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
+ if (q != Z3enums.INT_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ IntExpr(e)
+
+ let int_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ int_expr_of_arith_expr (arith_expr_of_expr (Expr.expr_of_ptr ctx no))
+
+ let int_num_of_int_expr e =
+ match e with IntExpr(ArithExpr(Expr(a))) ->
+ if (not (Z3native.is_numeral_ast (z3obj_gnc a) (z3obj_gno a))) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ IntNum(e)
+
+ let int_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ int_num_of_int_expr (int_expr_of_ptr ctx no)
+
+ let arith_sort_of_int_sort s = match s with IntSort(x) -> x
+ let arith_expr_of_int_expr e = match e with IntExpr(x) -> x
+ let int_expr_of_int_num e = match e with IntNum(x) -> x
+
+ let int_sort_of_arith_sort s = match s with ArithSort(Sort(a)) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.INT_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ IntSort(s)
+
+ let int_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ int_sort_of_arith_sort (arith_sort_of_sort (Sort.sort_of_ptr ctx no))
+
+ let sgc ( x : int_sort ) = match (x) with IntSort(s) -> (sgc s)
+ let sgnc ( x : int_sort ) = match (x) with IntSort(s) -> (sgnc s)
+ let sgno ( x : int_sort ) = match (x) with IntSort(s) -> (sgno s)
+ let egc ( x : int_expr ) = match (x) with IntExpr(e) -> (egc e)
+ let egnc ( x : int_expr ) = match (x) with IntExpr(e) -> (egnc e)
+ let egno ( x : int_expr ) = match (x) with IntExpr(e) -> (egno e)
+ let ngc ( x : int_num ) = match (x) with IntNum(e) -> (egc e)
+ let ngnc ( x : int_num ) = match (x) with IntNum(e) -> (egnc e)
+ let ngno ( x : int_num ) = match (x) with IntNum(e) -> (egno e)
+
+ let mk_sort ( ctx : context ) =
+ int_sort_of_ptr ctx (Z3native.mk_int_sort (context_gno ctx))
+
+ let get_int ( x : int_num ) =
+ let (r, v) = Z3native.get_numeral_int (ngnc x) (ngno x) in
+ if r then v
+ else raise (Z3native.Exception "Conversion failed.")
+
+ let to_string ( x : int_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+
+ let mk_int_const ( ctx : context ) ( name : Symbol.symbol ) =
+ IntExpr(ArithExpr(Expr.mk_const ctx name (match (mk_sort ctx) with IntSort(ArithSort(s)) -> s)))
+
+ let mk_int_const_s ( ctx : context ) ( name : string ) =
+ mk_int_const ctx (Symbol.mk_string ctx name)
+
+ let mk_mod ( ctx : context ) ( t1 : int_expr ) ( t2 : int_expr ) =
+ int_expr_of_ptr ctx (Z3native.mk_mod (context_gno ctx) (egno t1) (egno t2))
+
+ let mk_rem ( ctx : context ) ( t1 : int_expr ) ( t2 : int_expr ) =
+ int_expr_of_ptr ctx (Z3native.mk_rem (context_gno ctx) (egno t1) (egno t2))
+
+ let mk_int_numeral_s ( ctx : context ) ( v : string ) =
+ int_num_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (sgno (mk_sort ctx)))
+
+ let mk_int_numeral_i ( ctx : context ) ( v : int ) =
+ int_num_of_ptr ctx (Z3native.mk_int (context_gno ctx) v (sgno (mk_sort ctx)))
+
+ let mk_int2real ( ctx : context ) ( t : int_expr ) =
+ Real.real_expr_of_arith_expr (arith_expr_of_expr (Expr.expr_of_ptr ctx (Z3native.mk_int2real (context_gno ctx) (egno t))))
+
+ let mk_int2bv ( ctx : context ) ( n : int ) ( t : int_expr ) =
+ BitVector.bitvec_expr_of_expr (Expr.expr_of_ptr ctx (Z3native.mk_int2bv (context_gno ctx) n (egno t)))
+ end
+
+ and Real :
+ sig
+ type real_sort = RealSort of arith_sort
+ type real_expr = RealExpr of arith_expr
+ type rat_num = RatNum of real_expr
+
+ val real_expr_of_ptr : context -> Z3native.ptr -> real_expr
+ val rat_num_of_ptr : context -> Z3native.ptr -> rat_num
+
+ val arith_sort_of_real_sort : real_sort -> arith_sort
+ val real_sort_of_arith_sort : arith_sort -> real_sort
+ val arith_expr_of_real_expr : real_expr -> arith_expr
+ val real_expr_of_rat_num : rat_num -> real_expr
+ val real_expr_of_arith_expr : arith_expr -> real_expr
+ val rat_num_of_real_expr : real_expr -> rat_num
+
+ val mk_sort : context -> real_sort
+ val get_numerator : rat_num -> Integer.int_num
+ val get_denominator : rat_num -> Integer.int_num
+ val to_decimal_string : rat_num -> int -> string
+ val to_string : rat_num -> string
+ val mk_real_const : context -> Symbol.symbol -> real_expr
+ val mk_real_const_s : context -> string -> real_expr
+ val mk_numeral_nd : context -> int -> int -> rat_num
+ val mk_numeral_s : context -> string -> rat_num
+ val mk_numeral_i : context -> int -> rat_num
+ val mk_is_integer : context -> real_expr -> Boolean.bool_expr
+ val mk_real2int : context -> real_expr -> Integer.int_expr
+ end = struct
+ type real_sort = RealSort of arith_sort
+ type real_expr = RealExpr of arith_expr
+ type rat_num = RatNum of real_expr
+
+ let arith_sort_of_real_sort s = match s with RealSort(x) -> x
+ let arith_expr_of_real_expr e = match e with RealExpr(x) -> x
+ let real_expr_of_rat_num e = match e with RatNum(x) -> x
+
+ let real_expr_of_arith_expr e =
+ match e with ArithExpr(Expr(a)) ->
+ let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
+ if (q != Z3enums.REAL_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ RealExpr(e)
+
+ let real_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ real_expr_of_arith_expr (arith_expr_of_expr (Expr.expr_of_ptr ctx no))
+
+ let rat_num_of_real_expr e =
+ match e with RealExpr(ArithExpr(Expr(a))) ->
+ if (not (Z3native.is_numeral_ast (z3obj_gnc a) (z3obj_gno a))) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ RatNum(e)
+
+ let rat_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ rat_num_of_real_expr (real_expr_of_ptr ctx no)
+
+ let real_sort_of_arith_sort s = match s with ArithSort(Sort(a)) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.REAL_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ RealSort(s)
+
+ let real_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ real_sort_of_arith_sort (arith_sort_of_sort (sort_of_ptr ctx no))
+
+ let sgc ( x : real_sort ) = match (x) with RealSort(s) -> (sgc s)
+ let sgnc ( x : real_sort ) = match (x) with RealSort(s) -> (sgnc s)
+ let sgno ( x : real_sort ) = match (x) with RealSort(s) -> (sgno s)
+ let egc ( x : real_expr ) = match (x) with RealExpr(e) -> (egc e)
+ let egnc ( x : real_expr ) = match (x) with RealExpr(e) -> (egnc e)
+ let egno ( x : real_expr ) = match (x) with RealExpr(e) -> (egno e)
+ let ngc ( x : rat_num ) = match (x) with RatNum(e) -> (egc e)
+ let ngnc ( x : rat_num ) = match (x) with RatNum(e) -> (egnc e)
+ let ngno ( x : rat_num ) = match (x) with RatNum(e) -> (egno e)
+
+
+ let mk_sort ( ctx : context ) =
+ real_sort_of_ptr ctx (Z3native.mk_real_sort (context_gno ctx))
+
+ let get_numerator ( x : rat_num ) =
+ Integer.int_num_of_ptr (ngc x) (Z3native.get_numerator (ngnc x) (ngno x))
+
+ let get_denominator ( x : rat_num ) =
+ Integer.int_num_of_ptr (ngc x) (Z3native.get_denominator (ngnc x) (ngno x))
+
+ let to_decimal_string ( x : rat_num ) ( precision : int ) =
+ Z3native.get_numeral_decimal_string (ngnc x) (ngno x) precision
+
+ let to_string ( x : rat_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+
+ let mk_real_const ( ctx : context ) ( name : Symbol.symbol ) =
+ RealExpr(ArithExpr(Expr.mk_const ctx name (match (mk_sort ctx) with RealSort(ArithSort(s)) -> s)))
+
+ let mk_real_const_s ( ctx : context ) ( name : string ) =
+ mk_real_const ctx (Symbol.mk_string ctx name)
+
+ let mk_numeral_nd ( ctx : context ) ( num : int ) ( den : int) =
+ if (den == 0) then
+ raise (Z3native.Exception "Denominator is zero")
+ else
+ rat_num_of_ptr ctx (Z3native.mk_real (context_gno ctx) num den)
+
+ let mk_numeral_s ( ctx : context ) ( v : string ) =
+ rat_num_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (sgno (mk_sort ctx)))
+
+ let mk_numeral_i ( ctx : context ) ( v : int ) =
+ rat_num_of_ptr ctx (Z3native.mk_int (context_gno ctx) v (sgno (mk_sort ctx)))
+
+ let mk_is_integer ( ctx : context ) ( t : real_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_is_int (context_gno ctx) (egno t)))
+
+ let mk_real2int ( ctx : context ) ( t : real_expr ) =
+ Integer.int_expr_of_arith_expr (arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_real2int (context_gno ctx) (egno t))))
+ end
+
+ and AlgebraicNumber :
+ sig
+ type algebraic_num = AlgebraicNum of arith_expr
+
+ val arith_expr_of_algebraic_num : algebraic_num -> arith_expr
+ val algebraic_num_of_arith_expr : arith_expr -> algebraic_num
+
+ val to_upper : algebraic_num -> int -> Real.rat_num
+ val to_lower : algebraic_num -> int -> Real.rat_num
+ val to_decimal_string : algebraic_num -> int -> string
+ val to_string : algebraic_num -> string
+ end = struct
+ type algebraic_num = AlgebraicNum of arith_expr
+
+ let arith_expr_of_algebraic_num e = match e with AlgebraicNum(x) -> x
+
+ let algebraic_num_of_arith_expr e =
+ match e with ArithExpr(Expr(a)) ->
+ if (not (Z3native.is_algebraic_number (z3obj_gnc a) (z3obj_gno a))) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ AlgebraicNum(e)
+
+ let algebraic_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ algebraic_num_of_arith_expr (arith_expr_of_expr (expr_of_ptr ctx no))
+
+ let ngc ( x : algebraic_num ) = match (x) with AlgebraicNum(e) -> (egc e)
+ let ngnc ( x : algebraic_num ) = match (x) with AlgebraicNum(e) -> (egnc e)
+ let ngno ( x : algebraic_num ) = match (x) with AlgebraicNum(e) -> (egno e)
+
+
+ let to_upper ( x : algebraic_num ) ( precision : int ) =
+ Real.rat_num_of_ptr (ngc x) (Z3native.get_algebraic_number_upper (ngnc x) (ngno x) precision)
+
+ let to_lower ( x : algebraic_num ) precision =
+ Real.rat_num_of_ptr (ngc x) (Z3native.get_algebraic_number_lower (ngnc x) (ngno x) precision)
+
+ let to_decimal_string ( x : algebraic_num ) ( precision : int ) =
+ Z3native.get_numeral_decimal_string (ngnc x) (ngno x) precision
+
+ let to_string ( x : algebraic_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+ end
+
+ let is_int ( x : expr ) =
+ (Z3native.is_numeral_ast (Expr.gnc x) (Expr.gno x)) &&
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == INT_SORT)
+
+ let is_arithmetic_numeral ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ANUM)
+
+ let is_le ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_LE)
+
+ let is_ge ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_GE)
+
+ let is_lt ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_LT)
+
+ let is_gt ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_GT)
+
+ let is_add ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ADD)
+
+ let is_sub ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SUB)
+
+ let is_uminus ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_UMINUS)
+
+ let is_mul ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_MUL)
+
+ let is_div ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_DIV)
+
+ let is_idiv ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_IDIV)
+
+ let is_remainder ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_REM)
+
+ let is_modulus ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_MOD)
+
+ let is_inttoreal ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_TO_REAL)
+
+ let is_real_to_int ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_TO_INT)
+
+ let is_real_is_int ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_IS_INT)
+
+ let is_real ( x : expr ) =
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == REAL_SORT)
+ let is_int_numeral ( x : expr ) = (Expr.is_numeral x) && (is_int x)
+
+ let is_rat_num ( x : expr ) = (Expr.is_numeral x) && (is_real x)
+
+ let is_algebraic_number ( x : expr ) = Z3native.is_algebraic_number (Expr.gnc x) (Expr.gno x)
+
+ let mk_add ( ctx : context ) ( t : arith_expr list ) =
+ let f x = (Expr.gno (expr_of_arith_expr x)) in
+ arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_add (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
+
+ let mk_mul ( ctx : context ) ( t : arith_expr list ) =
+ let f x = (Expr.gno (expr_of_arith_expr x)) in
+ arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_mul (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
+
+ let mk_sub ( ctx : context ) ( t : arith_expr list ) =
+ let f x = (Expr.gno (expr_of_arith_expr x)) in
+ arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_sub (context_gno ctx) (List.length t) (Array.of_list (List.map f t))))
+
+ let mk_unary_minus ( ctx : context ) ( t : arith_expr ) =
+ arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_unary_minus (context_gno ctx) (egno t)))
+
+ let mk_div ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
+ arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_div (context_gno ctx) (egno t1) (egno t2)))
+
+ let mk_power ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
+ arith_expr_of_expr (expr_of_ptr ctx (Z3native.mk_power (context_gno ctx) (egno t1) (egno t2)))
+
+ let mk_lt ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_lt (context_gno ctx) (egno t1) (egno t2)))
+
+ let mk_le ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_le (context_gno ctx) (egno t1) (egno t2)))
+
+ let mk_gt ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_gt (context_gno ctx) (egno t1) (egno t2)))
+
+ let mk_ge ( ctx : context ) ( t1 : arith_expr ) ( t2 : arith_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_ge (context_gno ctx) (egno t1) (egno t2)))
+end
+
+
+and BitVector :
+sig
+ type bitvec_sort = BitVecSort of Sort.sort
+ type bitvec_expr = BitVecExpr of Expr.expr
+ type bitvec_num = BitVecNum of bitvec_expr
+
+ val sort_of_bitvec_sort : BitVector.bitvec_sort -> Sort.sort
+ val bitvec_sort_of_sort : Sort.sort -> BitVector.bitvec_sort
+ val expr_of_bitvec_expr : BitVector.bitvec_expr -> Expr.expr
+ val bitvec_expr_of_bitvec_num : BitVector.bitvec_num -> BitVector.bitvec_expr
+ val bitvec_expr_of_expr : Expr.expr -> BitVector.bitvec_expr
+ val bitvec_num_of_bitvec_expr : BitVector.bitvec_expr -> BitVector.bitvec_num
+
+ val mk_sort : context -> int -> bitvec_sort
+ val is_bv : Expr.expr -> bool
+ val is_bv_numeral : Expr.expr -> bool
+ val is_bv_bit1 : Expr.expr -> bool
+ val is_bv_bit0 : Expr.expr -> bool
+ val is_bv_uminus : Expr.expr -> bool
+ val is_bv_add : Expr.expr -> bool
+ val is_bv_sub : Expr.expr -> bool
+ val is_bv_mul : Expr.expr -> bool
+ val is_bv_sdiv : Expr.expr -> bool
+ val is_bv_udiv : Expr.expr -> bool
+ val is_bv_SRem : Expr.expr -> bool
+ val is_bv_urem : Expr.expr -> bool
+ val is_bv_smod : Expr.expr -> bool
+ val is_bv_sdiv0 : Expr.expr -> bool
+ val is_bv_udiv0 : Expr.expr -> bool
+ val is_bv_srem0 : Expr.expr -> bool
+ val is_bv_urem0 : Expr.expr -> bool
+ val is_bv_smod0 : Expr.expr -> bool
+ val is_bv_ule : Expr.expr -> bool
+ val is_bv_sle : Expr.expr -> bool
+ val is_bv_uge : Expr.expr -> bool
+ val is_bv_sge : Expr.expr -> bool
+ val is_bv_ult : Expr.expr -> bool
+ val is_bv_slt : Expr.expr -> bool
+ val is_bv_ugt : Expr.expr -> bool
+ val is_bv_sgt : Expr.expr -> bool
+ val is_bv_and : Expr.expr -> bool
+ val is_bv_or : Expr.expr -> bool
+ val is_bv_not : Expr.expr -> bool
+ val is_bv_xor : Expr.expr -> bool
+ val is_bv_nand : Expr.expr -> bool
+ val is_bv_nor : Expr.expr -> bool
+ val is_bv_xnor : Expr.expr -> bool
+ val is_bv_concat : Expr.expr -> bool
+ val is_bv_signextension : Expr.expr -> bool
+ val is_bv_zeroextension : Expr.expr -> bool
+ val is_bv_extract : Expr.expr -> bool
+ val is_bv_repeat : Expr.expr -> bool
+ val is_bv_reduceor : Expr.expr -> bool
+ val is_bv_reduceand : Expr.expr -> bool
+ val is_bv_comp : Expr.expr -> bool
+ val is_bv_shiftleft : Expr.expr -> bool
+ val is_bv_shiftrightlogical : Expr.expr -> bool
+ val is_bv_shiftrightarithmetic : Expr.expr -> bool
+ val is_bv_rotateleft : Expr.expr -> bool
+ val is_bv_rotateright : Expr.expr -> bool
+ val is_bv_rotateleftextended : Expr.expr -> bool
+ val is_bv_rotaterightextended : Expr.expr -> bool
+ val is_int_to_bv : Expr.expr -> bool
+ val is_bv_to_int : Expr.expr -> bool
+ val is_bv_carry : Expr.expr -> bool
+ val is_bv_xor3 : Expr.expr -> bool
+ val get_size : bitvec_sort -> int
+ val get_int : bitvec_num -> int
+ val to_string : bitvec_num -> string
+ val mk_const : context -> Symbol.symbol -> int -> bitvec_expr
+ val mk_const_s : context -> string -> int -> bitvec_expr
+ val mk_not : context -> bitvec_expr -> Expr.expr
+ val mk_redand : context -> bitvec_expr -> Expr.expr
+ val mk_redor : context -> bitvec_expr -> Expr.expr
+ val mk_and : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_or : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_xor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_nand : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_nor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_xnor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_neg : context -> bitvec_expr -> bitvec_expr
+ val mk_add : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_sub : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_mul : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_udiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_sdiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_urem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_srem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_smod : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_ult : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_slt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_ule : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sle : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_uge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_ugt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sgt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_concat : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_extract : context -> int -> int -> bitvec_expr -> bitvec_expr
+ val mk_sign_ext : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_zero_ext : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_repeat : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_shl : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_lshr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_ashr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_rotate_left : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_rotate_right : context -> int -> bitvec_expr -> bitvec_expr
+ val mk_ext_rotate_left : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_ext_rotate_right : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+ val mk_bv2int : context -> bitvec_expr -> bool -> Arithmetic.Integer.int_expr
+ val mk_add_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+ val mk_add_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sub_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_sub_no_underflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+ val mk_sdiv_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_neg_no_overflow : context -> bitvec_expr -> Boolean.bool_expr
+ val mk_mul_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+ val mk_mul_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+ val mk_numeral : context -> string -> int -> bitvec_num
+end = struct
+ type bitvec_sort = BitVecSort of sort
+ type bitvec_expr = BitVecExpr of expr
+ type bitvec_num = BitVecNum of bitvec_expr
+
+ let sort_of_bitvec_sort s = match s with BitVecSort(x) -> x
+
+ let bitvec_sort_of_sort s = match s with Sort(a) ->
+ if ((Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) (z3obj_gno a))) != Z3enums.BV_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ BitVecSort(s)
+
+ let bitvec_sort_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ bitvec_sort_of_sort (sort_of_ptr ctx no)
+
+ let bitvec_expr_of_expr e =
+ match e with Expr(a) ->
+ let s = Z3native.get_sort (z3obj_gnc a) (z3obj_gno a) in
+ let q = (Z3enums.sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc a) s)) in
+ if (q != Z3enums.BV_SORT) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ BitVecExpr(e)
+
+ let bitvec_expr_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ bitvec_expr_of_expr (expr_of_ptr ctx no)
+
+ let bitvec_num_of_bitvec_expr e =
+ match e with BitVecExpr(Expr(a)) ->
+ if (not (Z3native.is_numeral_ast (z3obj_gnc a) (z3obj_gno a))) then
+ raise (Z3native.Exception "Invalid coercion")
+ else
+ BitVecNum(e)
+
+ let bitvec_num_of_ptr ( ctx : context ) ( no : Z3native.ptr ) =
+ bitvec_num_of_bitvec_expr (bitvec_expr_of_expr (expr_of_ptr ctx no))
+
+ let expr_of_bitvec_expr e = match e with BitVecExpr(x) -> x
+ let bitvec_expr_of_bitvec_num e = match e with BitVecNum(x) -> x
+
+
+ let sgc ( x : bitvec_sort ) = match (x) with BitVecSort(s) -> (Sort.gc s)
+ let sgnc ( x : bitvec_sort ) = match (x) with BitVecSort(s) -> (Sort.gnc s)
+ let sgno ( x : bitvec_sort ) = match (x) with BitVecSort(s) -> (Sort.gno s)
+ let egc ( x : bitvec_expr ) = match (x) with BitVecExpr(e) -> (Expr.gc e)
+ let egnc ( x : bitvec_expr ) = match (x) with BitVecExpr(e) -> (Expr.gnc e)
+ let egno ( x : bitvec_expr ) = match (x) with BitVecExpr(e) -> (Expr.gno e)
+ let ngc ( x : bitvec_num ) = match (x) with BitVecNum(e) -> (egc e)
+ let ngnc ( x : bitvec_num ) = match (x) with BitVecNum(e) -> (egnc e)
+ let ngno ( x : bitvec_num ) = match (x) with BitVecNum(e) -> (egno e)
+
+
+ let mk_sort ( ctx : context ) size =
+ bitvec_sort_of_ptr ctx (Z3native.mk_bv_sort (context_gno ctx) size)
+ let is_bv ( x : expr ) =
+ ((sort_kind_of_int (Z3native.get_sort_kind (Expr.gnc x) (Z3native.get_sort (Expr.gnc x) (Expr.gno x)))) == BV_SORT)
+ let is_bv_numeral ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BNUM)
+ let is_bv_bit1 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BIT1)
+ let is_bv_bit0 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BIT0)
+ let is_bv_uminus ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BNEG)
+ let is_bv_add ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BADD)
+ let is_bv_sub ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSUB)
+ let is_bv_mul ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BMUL)
+ let is_bv_sdiv ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSDIV)
+ let is_bv_udiv ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BUDIV)
+ let is_bv_SRem ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSREM)
+ let is_bv_urem ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BUREM)
+ let is_bv_smod ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSMOD)
+ let is_bv_sdiv0 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSDIV0)
+ let is_bv_udiv0 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BUDIV0)
+ let is_bv_srem0 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSREM0)
+ let is_bv_urem0 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BUREM0)
+ let is_bv_smod0 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSMOD0)
+ let is_bv_ule ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ULEQ)
+ let is_bv_sle ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SLEQ)
+ let is_bv_uge ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_UGEQ)
+ let is_bv_sge ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SGEQ)
+ let is_bv_ult ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ULT)
+ let is_bv_slt ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SLT)
+ let is_bv_ugt ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_UGT)
+ let is_bv_sgt ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SGT)
+ let is_bv_and ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BAND)
+ let is_bv_or ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BOR)
+ let is_bv_not ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BNOT)
+ let is_bv_xor ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BXOR)
+ let is_bv_nand ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BNAND)
+ let is_bv_nor ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BNOR)
+ let is_bv_xnor ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BXNOR)
+ let is_bv_concat ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_CONCAT)
+ let is_bv_signextension ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_SIGN_EXT)
+ let is_bv_zeroextension ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ZERO_EXT)
+ let is_bv_extract ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_EXTRACT)
+ let is_bv_repeat ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_REPEAT)
+ let is_bv_reduceor ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BREDOR)
+ let is_bv_reduceand ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BREDAND)
+ let is_bv_comp ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BCOMP)
+ let is_bv_shiftleft ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BSHL)
+ let is_bv_shiftrightlogical ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BLSHR)
+ let is_bv_shiftrightarithmetic ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BASHR)
+ let is_bv_rotateleft ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ROTATE_LEFT)
+ let is_bv_rotateright ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_ROTATE_RIGHT)
+ let is_bv_rotateleftextended ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_EXT_ROTATE_LEFT)
+ let is_bv_rotaterightextended ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_EXT_ROTATE_RIGHT)
+ let is_int_to_bv ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_INT2BV)
+ let is_bv_to_int ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_BV2INT)
+ let is_bv_carry ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_CARRY)
+ let is_bv_xor3 ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_XOR3)
+ let get_size (x : bitvec_sort ) = Z3native.get_bv_sort_size (sgnc x) (sgno x)
+ let get_int ( x : bitvec_num ) =
+ let (r, v) = Z3native.get_numeral_int (ngnc x) (ngno x) in
+ if r then v
+ else raise (Z3native.Exception "Conversion failed.")
+ let to_string ( x : bitvec_num ) = Z3native.get_numeral_string (ngnc x) (ngno x)
+ let mk_const ( ctx : context ) ( name : Symbol.symbol ) ( size : int ) =
+ BitVecExpr(Expr.mk_const ctx name (match (BitVector.mk_sort ctx size) with BitVecSort(s) -> s))
+ let mk_const_s ( ctx : context ) ( name : string ) ( size : int ) =
+ mk_const ctx (Symbol.mk_string ctx name) size
+ let mk_not ( ctx : context ) ( t : bitvec_expr ) =
+ expr_of_ptr ctx (Z3native.mk_bvnot (context_gno ctx) (egno t))
+ let mk_redand ( ctx : context ) ( t : bitvec_expr) =
+ expr_of_ptr ctx (Z3native.mk_bvredand (context_gno ctx) (egno t))
+ let mk_redor ( ctx : context ) ( t : bitvec_expr) =
+ expr_of_ptr ctx (Z3native.mk_bvredor (context_gno ctx) (egno t))
+ let mk_and ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvand (context_gno ctx) (egno t1) (egno t2))
+ let mk_or ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvor (context_gno ctx) (egno t1) (egno t2))
+ let mk_xor ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvxor (context_gno ctx) (egno t1) (egno t2))
+ let mk_nand ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvnand (context_gno ctx) (egno t1) (egno t2))
+ let mk_nor ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvnor (context_gno ctx) (egno t1) (egno t2))
+ let mk_xnor ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvxnor (context_gno ctx) (egno t1) (egno t2))
+ let mk_neg ( ctx : context ) ( t : bitvec_expr) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvneg (context_gno ctx) (egno t))
+ let mk_add ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvadd (context_gno ctx) (egno t1) (egno t2))
+ let mk_sub ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvsub (context_gno ctx) (egno t1) (egno t2))
+ let mk_mul ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvmul (context_gno ctx) (egno t1) (egno t2))
+ let mk_udiv ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvudiv (context_gno ctx) (egno t1) (egno t2))
+ let mk_sdiv ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvsdiv (context_gno ctx) (egno t1) (egno t2))
+ let mk_urem ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvurem (context_gno ctx) (egno t1) (egno t2))
+ let mk_srem ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvsrem (context_gno ctx) (egno t1) (egno t2))
+ let mk_smod ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvsmod (context_gno ctx) (egno t1) (egno t2))
+ let mk_ult ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvult (context_gno ctx) (egno t1) (egno t2)))
+ let mk_slt ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvslt (context_gno ctx) (egno t1) (egno t2)))
+ let mk_ule ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvule (context_gno ctx) (egno t1) (egno t2)))
+ let mk_sle ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsle (context_gno ctx) (egno t1) (egno t2)))
+ let mk_uge ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvuge (context_gno ctx) (egno t1) (egno t2)))
+ let mk_sge ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsge (context_gno ctx) (egno t1) (egno t2)))
+ let mk_ugt ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvugt (context_gno ctx) (egno t1) (egno t2)))
+ let mk_sgt ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsgt (context_gno ctx) (egno t1) (egno t2)))
+ let mk_concat ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_concat (context_gno ctx) (egno t1) (egno t2))
+ let mk_extract ( ctx : context ) ( high : int ) ( low : int ) ( t : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_extract (context_gno ctx) high low (egno t))
+ let mk_sign_ext ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_sign_ext (context_gno ctx) i (egno t))
+ let mk_zero_ext ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_zero_ext (context_gno ctx) i (egno t))
+ let mk_repeat ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_repeat (context_gno ctx) i (egno t))
+ let mk_shl ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvshl (context_gno ctx) (egno t1) (egno t2))
+ let mk_lshr ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvlshr (context_gno ctx) (egno t1) (egno t2))
+ let mk_ashr ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_bvashr (context_gno ctx) (egno t1) (egno t2))
+ let mk_rotate_left ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_rotate_left (context_gno ctx) i (egno t))
+ let mk_rotate_right ( ctx : context ) ( i : int ) ( t : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_rotate_right (context_gno ctx) i (egno t))
+ let mk_ext_rotate_left ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_ext_rotate_left (context_gno ctx) (egno t1) (egno t2))
+ let mk_ext_rotate_right ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ bitvec_expr_of_ptr ctx (Z3native.mk_ext_rotate_right (context_gno ctx) (egno t1) (egno t2))
+ let mk_bv2int ( ctx : context ) ( t : bitvec_expr ) ( signed : bool ) =
+ Arithmetic.Integer.int_expr_of_ptr ctx (Z3native.mk_bv2int (context_gno ctx) (egno t) signed)
+ let mk_add_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) ( signed : bool) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvadd_no_overflow (context_gno ctx) (egno t1) (egno t2) signed))
+ let mk_add_no_underflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvadd_no_underflow (context_gno ctx) (egno t1) (egno t2)))
+ let mk_sub_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsub_no_overflow (context_gno ctx) (egno t1) (egno t2)))
+ let mk_sub_no_underflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) ( signed : bool) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsub_no_underflow (context_gno ctx) (egno t1) (egno t2) signed))
+ let mk_sdiv_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvsdiv_no_overflow (context_gno ctx) (egno t1) (egno t2)))
+ let mk_neg_no_overflow ( ctx : context ) ( t : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvneg_no_overflow (context_gno ctx) (egno t)))
+ let mk_mul_no_overflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) ( signed : bool) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvmul_no_overflow (context_gno ctx) (egno t1) (egno t2) signed))
+ let mk_mul_no_underflow ( ctx : context ) ( t1 : bitvec_expr ) ( t2 : bitvec_expr ) =
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.mk_bvmul_no_underflow (context_gno ctx) (egno t1) (egno t2)))
+ let mk_numeral ( ctx : context ) ( v : string ) ( size : int) =
+ bitvec_num_of_ptr ctx (Z3native.mk_numeral (context_gno ctx) v (sgno (BitVector.mk_sort ctx size)))
+end
+
+
+module Proof =
+struct
+ let is_true ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_TRUE)
+ let is_asserted ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_ASSERTED)
+ let is_goal ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_GOAL)
+ let is_modus_ponens ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_MODUS_PONENS)
+ let is_reflexivity ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_REFLEXIVITY)
+ let is_symmetry ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_SYMMETRY)
+ let is_transitivity ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_TRANSITIVITY)
+ let is_Transitivity_star ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_TRANSITIVITY_STAR)
+ let is_monotonicity ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_MONOTONICITY)
+ let is_quant_intro ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_QUANT_INTRO)
+ let is_distributivity ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_DISTRIBUTIVITY)
+ let is_and_elimination ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_AND_ELIM)
+ let is_or_elimination ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_NOT_OR_ELIM)
+ let is_rewrite ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_REWRITE)
+ let is_rewrite_star ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_REWRITE_STAR)
+ let is_pull_quant ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_PULL_QUANT)
+ let is_pull_quant_star ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_PULL_QUANT_STAR)
+ let is_push_quant ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_PUSH_QUANT)
+ let is_elim_unused_vars ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_ELIM_UNUSED_VARS)
+ let is_der ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_DER)
+ let is_quant_inst ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_QUANT_INST)
+ let is_hypothesis ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_HYPOTHESIS)
+ let is_lemma ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_LEMMA)
+ let is_unit_resolution ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_UNIT_RESOLUTION)
+ let is_iff_true ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_IFF_TRUE)
+ let is_iff_false ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_IFF_FALSE)
+ let is_commutativity ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_COMMUTATIVITY) (* *)
+ let is_def_axiom ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_DEF_AXIOM)
+ let is_def_intro ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_DEF_INTRO)
+ let is_apply_def ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_APPLY_DEF)
+ let is_iff_oeq ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_IFF_OEQ)
+ let is_nnf_pos ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_NNF_POS)
+ let is_nnf_neg ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_NNF_NEG)
+ let is_nnf_star ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_NNF_STAR)
+ let is_cnf_star ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_CNF_STAR)
+ let is_skolemize ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_SKOLEMIZE)
+ let is_modus_ponens_oeq ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_MODUS_PONENS_OEQ)
+ let is_theory_lemma ( x : expr ) = (FuncDecl.get_decl_kind (Expr.get_func_decl x) == OP_PR_TH_LEMMA)
+end
+
+
+module Goal =
+struct
+ type goal = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : goal = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.goal_inc_ref ;
+ dec_ref = Z3native.goal_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let get_precision ( x : goal ) =
+ goal_prec_of_int (Z3native.goal_precision (z3obj_gnc x) (z3obj_gno x))
+
+ let is_precise ( x : goal ) =
+ (get_precision x) == GOAL_PRECISE
+
+ let is_underapproximation ( x : goal ) =
+ (get_precision x) == GOAL_UNDER
+
+ let is_overapproximation ( x : goal ) =
+ (get_precision x) == GOAL_OVER
+
+ let is_garbage ( x : goal ) =
+ (get_precision x) == GOAL_UNDER_OVER
+
+ let assert_ ( x : goal ) ( constraints : Boolean.bool_expr list ) =
+ let f e = Z3native.goal_assert (z3obj_gnc x) (z3obj_gno x) (Boolean.gno e) in
+ ignore (List.map f constraints) ;
+ ()
+
+ let is_inconsistent ( x : goal ) =
+ Z3native.goal_inconsistent (z3obj_gnc x) (z3obj_gno x)
+
+ let get_depth ( x : goal ) = Z3native.goal_depth (z3obj_gnc x) (z3obj_gno x)
+
+ let reset ( x : goal ) = Z3native.goal_reset (z3obj_gnc x) (z3obj_gno x)
+
+ let get_size ( x : goal ) = Z3native.goal_size (z3obj_gnc x) (z3obj_gno x)
+
+ let get_formulas ( x : goal ) =
+ let n = get_size x in
+ let f i = (Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x)
+ (Z3native.goal_formula (z3obj_gnc x) (z3obj_gno x) i))) in
+ mk_list f n
+
+ let get_num_exprs ( x : goal ) = Z3native.goal_num_exprs (z3obj_gnc x) (z3obj_gno x)
+
+ let is_decided_sat ( x : goal ) =
+ Z3native.goal_is_decided_sat (z3obj_gnc x) (z3obj_gno x)
+
+ let is_decided_unsat ( x : goal ) =
+ Z3native.goal_is_decided_unsat (z3obj_gnc x) (z3obj_gno x)
+
+ let translate ( x : goal ) ( to_ctx : context ) =
+ create to_ctx (Z3native.goal_translate (z3obj_gnc x) (z3obj_gno x) (context_gno to_ctx))
+
+ let simplify ( x : goal ) ( p : Params.params option ) =
+ let tn = Z3native.mk_tactic (z3obj_gnc x) "simplify" in
+ Z3native.tactic_inc_ref (z3obj_gnc x) tn ;
+ let arn = match p with
+ | None -> Z3native.tactic_apply (z3obj_gnc x) tn (z3obj_gno x)
+ | Some(pn) -> Z3native.tactic_apply_ex (z3obj_gnc x) tn (z3obj_gno x) (z3obj_gno pn)
+ in
+ Z3native.apply_result_inc_ref (z3obj_gnc x) arn ;
+ let sg = Z3native.apply_result_get_num_subgoals (z3obj_gnc x) arn in
+ let res = if sg == 0 then
+ raise (Z3native.Exception "No subgoals")
+ else
+ Z3native.apply_result_get_subgoal (z3obj_gnc x) arn 0 in
+ Z3native.apply_result_dec_ref (z3obj_gnc x) arn ;
+ Z3native.tactic_dec_ref (z3obj_gnc x) tn ;
+ create (z3obj_gc x) res
+
+ let mk_goal ( ctx : context ) ( models : bool ) ( unsat_cores : bool ) ( proofs : bool ) =
+ create ctx (Z3native.mk_goal (context_gno ctx) models unsat_cores proofs)
+
+ let to_string ( x : goal ) = Z3native.goal_to_string (z3obj_gnc x) (z3obj_gno x)
+end
+
+
+module Model =
+struct
+ type model = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : model = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.model_inc_ref ;
+ dec_ref = Z3native.model_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ module FuncInterp =
+ struct
+ type func_interp = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : func_interp = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.func_interp_inc_ref ;
+ dec_ref = Z3native.func_interp_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ module FuncEntry =
+ struct
+ type func_entry = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : func_entry = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.func_entry_inc_ref ;
+ dec_ref = Z3native.func_entry_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let get_value ( x : func_entry ) =
+ expr_of_ptr (z3obj_gc x) (Z3native.func_entry_get_value (z3obj_gnc x) (z3obj_gno x))
+
+ let get_num_args ( x : func_entry ) = Z3native.func_entry_get_num_args (z3obj_gnc x) (z3obj_gno x)
+
+ let get_args ( x : func_entry ) =
+ let n = (get_num_args x) in
+ let f i = (expr_of_ptr (z3obj_gc x) (Z3native.func_entry_get_arg (z3obj_gnc x) (z3obj_gno x) i)) in
+ mk_list f n
+
+ let to_string ( x : func_entry ) =
+ let a = (get_args x) in
+ let f c p = (p ^ (Expr.to_string c) ^ ", ") in
+ "[" ^ List.fold_right f a ((Expr.to_string (get_value x)) ^ "]")
+ end
+
+ let get_num_entries ( x: func_interp ) = Z3native.func_interp_get_num_entries (z3obj_gnc x) (z3obj_gno x)
+
+ let get_entries ( x : func_interp ) =
+ let n = (get_num_entries x) in
+ let f i = (FuncEntry.create (z3obj_gc x) (Z3native.func_interp_get_entry (z3obj_gnc x) (z3obj_gno x) i)) in
+ mk_list f n
+
+ let get_else ( x : func_interp ) = expr_of_ptr (z3obj_gc x) (Z3native.func_interp_get_else (z3obj_gnc x) (z3obj_gno x))
+
+ let get_arity ( x : func_interp ) = Z3native.func_interp_get_arity (z3obj_gnc x) (z3obj_gno x)
+
+ let to_string ( x : func_interp ) =
+ let f c p = (
+ let n = (FuncEntry.get_num_args c) in
+ p ^
+ let g c p = (p ^ (Expr.to_string c) ^ ", ") in
+ (if n > 1 then "[" else "") ^
+ (List.fold_right
+ g
+ (FuncEntry.get_args c)
+ ((if n > 1 then "]" else "") ^ " -> " ^ (Expr.to_string (FuncEntry.get_value c)) ^ ", "))
+ ) in
+ List.fold_right f (get_entries x) ("else -> " ^ (Expr.to_string (get_else x)) ^ "]")
+ end
+
+ let get_const_interp ( x : model ) ( f : func_decl ) =
+ if (FuncDecl.get_arity f) != 0 ||
+ (sort_kind_of_int (Z3native.get_sort_kind (FuncDecl.gnc f) (Z3native.get_range (FuncDecl.gnc f) (FuncDecl.gno f)))) == ARRAY_SORT then
+ raise (Z3native.Exception "Non-zero arity functions and arrays have FunctionInterpretations as a model. Use FuncInterp.")
+ else
+ let np = Z3native.model_get_const_interp (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno f) in
+ if (Z3native.is_null np) then
+ None
+ else
+ Some (expr_of_ptr (z3obj_gc x) np)
+
+ let get_const_interp_e ( x : model ) ( a : expr ) = get_const_interp x (Expr.get_func_decl a)
+
+
+ let rec get_func_interp ( x : model ) ( f : func_decl ) =
+ let sk = (sort_kind_of_int (Z3native.get_sort_kind (z3obj_gnc x) (Z3native.get_range (FuncDecl.gnc f) (FuncDecl.gno f)))) in
+ if (FuncDecl.get_arity f) == 0 then
+ let n = Z3native.model_get_const_interp (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno f) in
+ if (Z3native.is_null n) then
+ None
+ else
+ match sk with
+ | ARRAY_SORT ->
+ if not (Z3native.is_as_array (z3obj_gnc x) n) then
+ raise (Z3native.Exception "Argument was not an array constant")
+ else
+ let fd = Z3native.get_as_array_func_decl (z3obj_gnc x) n in
+ get_func_interp x (func_decl_of_ptr (z3obj_gc x) fd)
+ | _ -> raise (Z3native.Exception "Constant functions do not have a function interpretation; use ConstInterp");
+ else
+ let n = (Z3native.model_get_func_interp (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno f)) in
+ if (Z3native.is_null n) then None else Some (FuncInterp.create (z3obj_gc x) n)
+
+ (** The number of constants that have an interpretation in the model. *)
+ let get_num_consts ( x : model ) = Z3native.model_get_num_consts (z3obj_gnc x) (z3obj_gno x)
+
+ let get_const_decls ( x : model ) =
+ let n = (get_num_consts x) in
+ let f i = func_decl_of_ptr (z3obj_gc x) (Z3native.model_get_const_decl (z3obj_gnc x) (z3obj_gno x) i) in
+ mk_list f n
+
+ let get_num_funcs ( x : model ) = Z3native.model_get_num_funcs (z3obj_gnc x) (z3obj_gno x)
+
+ let get_func_decls ( x : model ) =
+ let n = (get_num_consts x) in
+ let f i = func_decl_of_ptr (z3obj_gc x) (Z3native.model_get_func_decl (z3obj_gnc x) (z3obj_gno x) i) in
+ mk_list f n
+
+ let get_decls ( x : model ) =
+ let n_funcs = (get_num_funcs x) in
+ let n_consts = (get_num_consts x ) in
+ let f i = func_decl_of_ptr (z3obj_gc x) (Z3native.model_get_func_decl (z3obj_gnc x) (z3obj_gno x) i) in
+ let g i = func_decl_of_ptr (z3obj_gc x) (Z3native.model_get_const_decl (z3obj_gnc x) (z3obj_gno x) i) in
+ (mk_list f n_funcs) @ (mk_list g n_consts)
+
+ exception ModelEvaluationFailedException of string
+
+ let eval ( x : model ) ( t : expr ) ( completion : bool ) =
+ let (r, v) = (Z3native.model_eval (z3obj_gnc x) (z3obj_gno x) (Expr.gno t) completion) in
+ if not r then
+ raise (ModelEvaluationFailedException "evaluation failed")
+ else
+ expr_of_ptr (z3obj_gc x) v
+
+ let evaluate ( x : model ) ( t : expr ) ( completion : bool ) =
+ eval x t completion
+
+ let get_num_sorts ( x : model ) = Z3native.model_get_num_sorts (z3obj_gnc x) (z3obj_gno x)
+
+ let get_sorts ( x : model ) =
+ let n = (get_num_sorts x) in
+ let f i = (sort_of_ptr (z3obj_gc x) (Z3native.model_get_sort (z3obj_gnc x) (z3obj_gno x) i)) in
+ mk_list f n
+
+ let sort_universe ( x : model ) ( s : sort ) =
+ let n_univ = AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.model_get_sort_universe (z3obj_gnc x) (z3obj_gno x) (Sort.gno s)) in
+ let n = (AST.ASTVector.get_size n_univ) in
+ let f i = (AST.ASTVector.get n_univ i) in
+ mk_list f n
+
+ let to_string ( x : model ) = Z3native.model_to_string (z3obj_gnc x) (z3obj_gno x)
+end
+
+
+module Probe =
+struct
+ type probe = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : probe = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.probe_inc_ref ;
+ dec_ref = Z3native.probe_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+
+ let apply ( x : probe ) ( g : Goal.goal ) =
+ Z3native.probe_apply (z3obj_gnc x) (z3obj_gno x) (z3obj_gno g)
+
+ let get_num_probes ( ctx : context ) =
+ Z3native.get_num_probes (context_gno ctx)
+
+ let get_probe_names ( ctx : context ) =
+ let n = (get_num_probes ctx) in
+ let f i = (Z3native.get_probe_name (context_gno ctx) i) in
+ mk_list f n
+
+ let get_probe_description ( ctx : context ) ( name : string ) =
+ Z3native.probe_get_descr (context_gno ctx) name
+
+ let mk_probe ( ctx : context ) ( name : string ) =
+ (create ctx (Z3native.mk_probe (context_gno ctx) name))
+
+ let const ( ctx : context ) ( v : float ) =
+ (create ctx (Z3native.probe_const (context_gno ctx) v))
+
+ let lt ( ctx : context ) ( p1 : probe ) ( p2 : probe ) =
+ (create ctx (Z3native.probe_lt (context_gno ctx) (z3obj_gno p1) (z3obj_gno p2)))
+
+ let gt ( ctx : context ) ( p1 : probe ) ( p2 : probe ) =
+ (create ctx (Z3native.probe_gt (context_gno ctx) (z3obj_gno p1) (z3obj_gno p2)))
+
+ let le ( ctx : context ) ( p1 : probe ) ( p2 : probe ) =
+ (create ctx (Z3native.probe_le (context_gno ctx) (z3obj_gno p1) (z3obj_gno p2)))
+
+ let ge ( ctx : context ) ( p1 : probe ) ( p2 : probe ) =
+ (create ctx (Z3native.probe_ge (context_gno ctx) (z3obj_gno p1) (z3obj_gno p2)))
+
+ let eq ( ctx : context ) ( p1 : probe ) ( p2 : probe ) =
+ (create ctx (Z3native.probe_eq (context_gno ctx) (z3obj_gno p1) (z3obj_gno p2)))
+
+ let and_ ( ctx : context ) ( p1 : probe ) ( p2 : probe ) =
+ (create ctx (Z3native.probe_and (context_gno ctx) (z3obj_gno p1) (z3obj_gno p2)))
+
+ let or_ ( ctx : context ) ( p1 : probe ) ( p2 : probe ) =
+ (create ctx (Z3native.probe_or (context_gno ctx) (z3obj_gno p1) (z3obj_gno p2)))
+
+ let not_ ( ctx : context ) ( p : probe ) =
+ (create ctx (Z3native.probe_not (context_gno ctx) (z3obj_gno p)))
+end
+
+
+module Tactic =
+struct
+ type tactic = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : tactic = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.tactic_inc_ref ;
+ dec_ref = Z3native.tactic_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ module ApplyResult =
+ struct
+ type apply_result = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : apply_result = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.apply_result_inc_ref ;
+ dec_ref = Z3native.apply_result_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let get_num_subgoals ( x : apply_result ) =
+ Z3native.apply_result_get_num_subgoals (z3obj_gnc x) (z3obj_gno x)
+
+ let get_subgoals ( x : apply_result ) =
+ let n = (get_num_subgoals x) in
+ let f i = Goal.create (z3obj_gc x) (Z3native.apply_result_get_subgoal (z3obj_gnc x) (z3obj_gno x) i) in
+ mk_list f n
+
+ let get_subgoal ( x : apply_result ) ( i : int ) =
+ Goal.create (z3obj_gc x) (Z3native.apply_result_get_subgoal (z3obj_gnc x) (z3obj_gno x) i)
+
+ let convert_model ( x : apply_result ) ( i : int ) ( m : Model.model ) =
+ Model.create (z3obj_gc x) (Z3native.apply_result_convert_model (z3obj_gnc x) (z3obj_gno x) i (z3obj_gno m))
+
+ let to_string ( x : apply_result ) = Z3native.apply_result_to_string (z3obj_gnc x) (z3obj_gno x)
+ end
+
+ let get_help ( x : tactic ) = Z3native.tactic_get_help (z3obj_gnc x) (z3obj_gno x)
+
+ let get_param_descrs ( x : tactic ) =
+ Params.ParamDescrs.param_descrs_of_ptr (z3obj_gc x) (Z3native.tactic_get_param_descrs (z3obj_gnc x) (z3obj_gno x))
+
+ let apply ( x : tactic ) ( g : Goal.goal ) ( p : Params.params option ) =
+ match p with
+ | None -> (ApplyResult.create (z3obj_gc x) (Z3native.tactic_apply (z3obj_gnc x) (z3obj_gno x) (z3obj_gno g)))
+ | Some (pn) -> (ApplyResult.create (z3obj_gc x) (Z3native.tactic_apply_ex (z3obj_gnc x) (z3obj_gno x) (z3obj_gno g) (z3obj_gno pn)))
+
+ let get_num_tactics ( ctx : context ) = Z3native.get_num_tactics (context_gno ctx)
+
+ let get_tactic_names ( ctx : context ) =
+ let n = (get_num_tactics ctx ) in
+ let f i = (Z3native.get_tactic_name (context_gno ctx) i) in
+ mk_list f n
+
+ let get_tactic_description ( ctx : context ) ( name : string ) =
+ Z3native.tactic_get_descr (context_gno ctx) name
+
+ let mk_tactic ( ctx : context ) ( name : string ) =
+ create ctx (Z3native.mk_tactic (context_gno ctx) name)
+
+ let and_then ( ctx : context ) ( t1 : tactic ) ( t2 : tactic ) ( ts : tactic list ) =
+ let f p c = (match p with
+ | None -> (Some (z3obj_gno c))
+ | Some(x) -> (Some (Z3native.tactic_and_then (context_gno ctx) (z3obj_gno c) x))) in
+ match (List.fold_left f None ts) with
+ | None ->
+ create ctx (Z3native.tactic_and_then (context_gno ctx) (z3obj_gno t1) (z3obj_gno t2))
+ | Some(x) ->
+ let o = (Z3native.tactic_and_then (context_gno ctx) (z3obj_gno t2) x) in
+ create ctx (Z3native.tactic_and_then (context_gno ctx) (z3obj_gno t1) o)
+
+ let or_else ( ctx : context ) ( t1 : tactic ) ( t2 : tactic ) =
+ create ctx (Z3native.tactic_or_else (context_gno ctx) (z3obj_gno t1) (z3obj_gno t2))
+
+ let try_for ( ctx : context ) ( t : tactic ) ( ms : int ) =
+ create ctx (Z3native.tactic_try_for (context_gno ctx) (z3obj_gno t) ms)
+
+ let when_ ( ctx : context ) ( p : Probe.probe ) ( t : tactic ) =
+ create ctx (Z3native.tactic_when (context_gno ctx) (z3obj_gno p) (z3obj_gno t))
+
+ let cond ( ctx : context ) ( p : Probe.probe ) ( t1 : tactic ) ( t2 : tactic ) =
+ create ctx (Z3native.tactic_cond (context_gno ctx) (z3obj_gno p) (z3obj_gno t1) (z3obj_gno t2))
+
+ let repeat ( ctx : context ) ( t : tactic ) ( max : int ) =
+ create ctx (Z3native.tactic_repeat (context_gno ctx) (z3obj_gno t) max)
+
+ let skip ( ctx : context ) =
+ create ctx (Z3native.tactic_skip (context_gno ctx))
+
+ let fail ( ctx : context ) =
+ create ctx (Z3native.tactic_fail (context_gno ctx))
+
+ let fail_if ( ctx : context ) ( p : Probe.probe ) =
+ create ctx (Z3native.tactic_fail_if (context_gno ctx) (z3obj_gno p))
+
+ let fail_if_not_decided ( ctx : context ) =
+ create ctx (Z3native.tactic_fail_if_not_decided (context_gno ctx))
+
+ let using_params ( ctx : context ) ( t : tactic ) ( p : Params.params ) =
+ create ctx (Z3native.tactic_using_params (context_gno ctx) (z3obj_gno t) (z3obj_gno p))
+
+ let with_ ( ctx : context ) ( t : tactic ) ( p : Params.params ) =
+ using_params ctx t p
+
+ let par_or ( ctx : context ) ( t : tactic list ) =
+ let f e = (z3obj_gno e) in
+ create ctx (Z3native.tactic_par_or (context_gno ctx) (List.length t) (Array.of_list (List.map f t)))
+
+ let par_and_then ( ctx : context ) ( t1 : tactic ) ( t2 : tactic ) =
+ create ctx (Z3native.tactic_par_and_then (context_gno ctx) (z3obj_gno t1) (z3obj_gno t2))
+
+ let interrupt ( ctx : context ) =
+ Z3native.interrupt (context_gno ctx)
+end
+
+
+module Solver =
+struct
+ type solver = z3_native_object
+ type status = UNSATISFIABLE | UNKNOWN | SATISFIABLE
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : solver = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.solver_inc_ref ;
+ dec_ref = Z3native.solver_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+ let string_of_status ( s : status) = match s with
+ | UNSATISFIABLE -> "unsatisfiable"
+ | SATISFIABLE -> "satisfiable"
+ | _ -> "unknown"
+
+ module Statistics =
+ struct
+ type statistics = z3_native_object
+
+ let create ( ctx : context ) ( no : Z3native.ptr ) =
+ let res : statistics = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.stats_inc_ref ;
+ dec_ref = Z3native.stats_dec_ref } in
+ (z3obj_sno res ctx no) ;
+ (z3obj_create res) ;
+ res
+
+
+ module Entry =
+ struct
+ type statistics_entry = {
+ mutable m_key : string;
+ mutable m_is_int : bool ;
+ mutable m_is_float : bool ;
+ mutable m_int : int ;
+ mutable m_float : float }
+
+ let create_si k v =
+ let res : statistics_entry = {
+ m_key = k ;
+ m_is_int = true ;
+ m_is_float = false ;
+ m_int = v ;
+ m_float = 0.0
+ } in
+ res
+
+ let create_sd k v =
+ let res : statistics_entry = {
+ m_key = k ;
+ m_is_int = false ;
+ m_is_float = true ;
+ m_int = 0 ;
+ m_float = v
+ } in
+ res
+
+
+ let get_key (x : statistics_entry) = x.m_key
+ let get_int (x : statistics_entry) = x.m_int
+ let get_float (x : statistics_entry) = x.m_float
+ let is_int (x : statistics_entry) = x.m_is_int
+ let is_float (x : statistics_entry) = x.m_is_float
+ let to_string_value (x : statistics_entry) =
+ if (is_int x) then
+ string_of_int (get_int x)
+ else if (is_float x) then
+ string_of_float (get_float x)
+ else
+ raise (Z3native.Exception "Unknown statistical entry type")
+ let to_string ( x : statistics_entry ) = (get_key x) ^ ": " ^ (to_string_value x)
+ end
+
+ let to_string ( x : statistics ) = Z3native.stats_to_string (z3obj_gnc x) (z3obj_gno x)
+
+ let get_size ( x : statistics ) = Z3native.stats_size (z3obj_gnc x) (z3obj_gno x)
+
+ let get_entries ( x : statistics ) =
+ let n = (get_size x ) in
+ let f i = (
+ let k = Z3native.stats_get_key (z3obj_gnc x) (z3obj_gno x) i in
+ if (Z3native.stats_is_uint (z3obj_gnc x) (z3obj_gno x) i) then
+ (Entry.create_si k (Z3native.stats_get_uint_value (z3obj_gnc x) (z3obj_gno x) i))
+ else
+ (Entry.create_sd k (Z3native.stats_get_double_value (z3obj_gnc x) (z3obj_gno x) i))
+ ) in
+ mk_list f n
+
+ let get_keys ( x : statistics ) =
+ let n = (get_size x) in
+ let f i = (Z3native.stats_get_key (z3obj_gnc x) (z3obj_gno x) i) in
+ mk_list f n
+
+ let get ( x : statistics ) ( key : string ) =
+ let f p c = (if ((Entry.get_key c) == key) then (Some c) else p) in
+ List.fold_left f None (get_entries x)
+ end
+
+ let get_help ( x : solver ) = Z3native.solver_get_help (z3obj_gnc x) (z3obj_gno x)
+
+ let set_parameters ( x : solver ) ( p : Params.params )=
+ Z3native.solver_set_params (z3obj_gnc x) (z3obj_gno x) (z3obj_gno p)
+
+ let get_param_descrs ( x : solver ) =
+ Params.ParamDescrs.param_descrs_of_ptr (z3obj_gc x) (Z3native.solver_get_param_descrs (z3obj_gnc x) (z3obj_gno x))
+
+ let get_num_scopes ( x : solver ) = Z3native.solver_get_num_scopes (z3obj_gnc x) (z3obj_gno x)
+
+ let push ( x : solver ) = Z3native.solver_push (z3obj_gnc x) (z3obj_gno x)
+
+ let pop ( x : solver ) ( n : int ) = Z3native.solver_pop (z3obj_gnc x) (z3obj_gno x) n
+
+ let reset ( x : solver ) = Z3native.solver_reset (z3obj_gnc x) (z3obj_gno x)
+
+ let assert_ ( x : solver ) ( constraints : Boolean.bool_expr list ) =
+ let f e = (Z3native.solver_assert (z3obj_gnc x) (z3obj_gno x) (Boolean.gno e)) in
+ ignore (List.map f constraints)
+
+ let assert_and_track_a ( x : solver ) ( cs : Boolean.bool_expr list ) ( ps : Boolean.bool_expr list ) =
+ if ((List.length cs) != (List.length ps)) then
+ raise (Z3native.Exception "Argument size mismatch")
+ else
+ let f a b = (Z3native.solver_assert_and_track (z3obj_gnc x) (z3obj_gno x) (Boolean.gno a) (Boolean.gno b)) in
+ ignore (List.iter2 f cs ps)
+
+ let assert_and_track ( x : solver ) ( c : Boolean.bool_expr ) ( p : Boolean.bool_expr ) =
+ Z3native.solver_assert_and_track (z3obj_gnc x) (z3obj_gno x) (Boolean.gno c) (Boolean.gno p)
+
+ let get_num_assertions ( x : solver ) =
+ let a = AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.solver_get_assertions (z3obj_gnc x) (z3obj_gno x)) in
+ (AST.ASTVector.get_size a)
+
+ let get_assertions ( x : solver ) =
+ let a = AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.solver_get_assertions (z3obj_gnc x) (z3obj_gno x)) in
+ let n = (AST.ASTVector.get_size a) in
+ let f i = Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get a i))) in
+ mk_list f n
+
+ let check ( x : solver ) ( assumptions : Boolean.bool_expr list ) =
+ let r =
+ if ((List.length assumptions) == 0) then
+ lbool_of_int (Z3native.solver_check (z3obj_gnc x) (z3obj_gno x))
+ else
+ let f x = (Expr.gno (Boolean.expr_of_bool_expr x)) in
+ lbool_of_int (Z3native.solver_check_assumptions (z3obj_gnc x) (z3obj_gno x) (List.length assumptions) (Array.of_list (List.map f assumptions)))
+ in
+ match r with
+ | L_TRUE -> SATISFIABLE
+ | L_FALSE -> UNSATISFIABLE
+ | _ -> UNKNOWN
+
+ let get_model ( x : solver ) =
+ let q = Z3native.solver_get_model (z3obj_gnc x) (z3obj_gno x) in
+ if (Z3native.is_null q) then
+ None
+ else
+ Some (Model.create (z3obj_gc x) q)
+
+ let get_proof ( x : solver ) =
+ let q = Z3native.solver_get_proof (z3obj_gnc x) (z3obj_gno x) in
+ if (Z3native.is_null q) then
+ None
+ else
+ Some (expr_of_ptr (z3obj_gc x) q)
+
+ let get_unsat_core ( x : solver ) =
+ let cn = AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.solver_get_unsat_core (z3obj_gnc x) (z3obj_gno x)) in
+ let n = (AST.ASTVector.get_size cn) in
+ let f i = (AST.ASTVector.get cn i) in
+ mk_list f n
+
+ let get_reason_unknown ( x : solver ) = Z3native.solver_get_reason_unknown (z3obj_gnc x) (z3obj_gno x)
+
+ let get_statistics ( x : solver ) =
+ (Statistics.create (z3obj_gc x) (Z3native.solver_get_statistics (z3obj_gnc x) (z3obj_gno x)))
+
+ let mk_solver ( ctx : context ) ( logic : Symbol.symbol option ) =
+ match logic with
+ | None -> (create ctx (Z3native.mk_solver (context_gno ctx)))
+ | Some (x) -> (create ctx (Z3native.mk_solver_for_logic (context_gno ctx) (Symbol.gno x)))
+
+ let mk_solver_s ( ctx : context ) ( logic : string ) =
+ mk_solver ctx (Some (Symbol.mk_string ctx logic))
+
+ let mk_simple_solver ( ctx : context ) =
+ (create ctx (Z3native.mk_simple_solver (context_gno ctx)))
+
+ let mk_solver_t ( ctx : context ) ( t : Tactic.tactic ) =
+ (create ctx (Z3native.mk_solver_from_tactic (context_gno ctx) (z3obj_gno t)))
+
+ let to_string ( x : solver ) = Z3native.solver_to_string (z3obj_gnc x) (z3obj_gno x)
+end
+
+
+module Fixedpoint =
+struct
+ type fixedpoint = z3_native_object
+
+ let create ( ctx : context ) =
+ let res : fixedpoint = { m_ctx = ctx ;
+ m_n_obj = null ;
+ inc_ref = Z3native.fixedpoint_inc_ref ;
+ dec_ref = Z3native.fixedpoint_dec_ref } in
+ (z3obj_sno res ctx (Z3native.mk_fixedpoint (context_gno ctx))) ;
+ (z3obj_create res) ;
+ res
+
+
+ let get_help ( x : fixedpoint ) =
+ Z3native.fixedpoint_get_help (z3obj_gnc x) (z3obj_gno x)
+
+ let set_params ( x : fixedpoint ) ( p : Params.params )=
+ Z3native.fixedpoint_set_params (z3obj_gnc x) (z3obj_gno x) (z3obj_gno p)
+
+ let get_param_descrs ( x : fixedpoint ) =
+ Params.ParamDescrs.param_descrs_of_ptr (z3obj_gc x) (Z3native.fixedpoint_get_param_descrs (z3obj_gnc x) (z3obj_gno x))
+
+ let assert_ ( x : fixedpoint ) ( constraints : Boolean.bool_expr list ) =
+ let f e = (Z3native.fixedpoint_assert (z3obj_gnc x) (z3obj_gno x) (Boolean.gno e)) in
+ ignore (List.map f constraints) ;
+ ()
+
+ let register_relation ( x : fixedpoint ) ( f : func_decl ) =
+ Z3native.fixedpoint_register_relation (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno f)
+
+ let add_rule ( x : fixedpoint ) ( rule : Boolean.bool_expr ) ( name : Symbol.symbol option ) =
+ match name with
+ | None -> Z3native.fixedpoint_add_rule (z3obj_gnc x) (z3obj_gno x) (Boolean.gno rule) null
+ | Some(y) -> Z3native.fixedpoint_add_rule (z3obj_gnc x) (z3obj_gno x) (Boolean.gno rule) (Symbol.gno y)
+
+ let add_fact ( x : fixedpoint ) ( pred : func_decl ) ( args : int list ) =
+ Z3native.fixedpoint_add_fact (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno pred) (List.length args) (Array.of_list args)
+
+ let query ( x : fixedpoint ) ( query : Boolean.bool_expr ) =
+ match (lbool_of_int (Z3native.fixedpoint_query (z3obj_gnc x) (z3obj_gno x) (Boolean.gno query))) with
+ | L_TRUE -> Solver.SATISFIABLE
+ | L_FALSE -> Solver.UNSATISFIABLE
+ | _ -> Solver.UNKNOWN
+
+ let query_r ( x : fixedpoint ) ( relations : func_decl list ) =
+ let f x = ptr_of_ast (ast_of_func_decl x) in
+ match (lbool_of_int (Z3native.fixedpoint_query_relations (z3obj_gnc x) (z3obj_gno x) (List.length relations) (Array.of_list (List.map f relations)))) with
+ | L_TRUE -> Solver.SATISFIABLE
+ | L_FALSE -> Solver.UNSATISFIABLE
+ | _ -> Solver.UNKNOWN
+
+ let push ( x : fixedpoint ) =
+ Z3native.fixedpoint_push (z3obj_gnc x) (z3obj_gno x)
+
+ let pop ( x : fixedpoint ) =
+ Z3native.fixedpoint_pop (z3obj_gnc x) (z3obj_gno x)
+
+ let update_rule ( x : fixedpoint ) ( rule : Boolean.bool_expr ) ( name : Symbol.symbol ) =
+ Z3native.fixedpoint_update_rule (z3obj_gnc x) (z3obj_gno x) (Boolean.gno rule) (Symbol.gno name)
+
+ let get_answer ( x : fixedpoint ) =
+ let q = (Z3native.fixedpoint_get_answer (z3obj_gnc x) (z3obj_gno x)) in
+ if (Z3native.is_null q) then
+ None
+ else
+ Some (expr_of_ptr (z3obj_gc x) q)
+
+ let get_reason_unknown ( x : fixedpoint ) =
+ Z3native.fixedpoint_get_reason_unknown (z3obj_gnc x) (z3obj_gno x)
+
+ let get_num_levels ( x : fixedpoint ) ( predicate : func_decl ) =
+ Z3native.fixedpoint_get_num_levels (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno predicate)
+
+ let get_cover_delta ( x : fixedpoint ) ( level : int ) ( predicate : func_decl ) =
+ let q = (Z3native.fixedpoint_get_cover_delta (z3obj_gnc x) (z3obj_gno x) level (FuncDecl.gno predicate)) in
+ if (Z3native.is_null q) then
+ None
+ else
+ Some (expr_of_ptr (z3obj_gc x) q)
+
+ let add_cover ( x : fixedpoint ) ( level : int ) ( predicate : func_decl ) ( property : expr ) =
+ Z3native.fixedpoint_add_cover (z3obj_gnc x) (z3obj_gno x) level (FuncDecl.gno predicate) (Expr.gno property)
+
+ let to_string ( x : fixedpoint ) = Z3native.fixedpoint_to_string (z3obj_gnc x) (z3obj_gno x) 0 [||]
+
+ let set_predicate_representation ( x : fixedpoint ) ( f : func_decl ) ( kinds : Symbol.symbol list ) =
+ Z3native.fixedpoint_set_predicate_representation (z3obj_gnc x) (z3obj_gno x) (FuncDecl.gno f) (List.length kinds) (Symbol.symbol_lton kinds)
+
+ let to_string_q ( x : fixedpoint ) ( queries : Boolean.bool_expr list ) =
+ let f x = ptr_of_expr (Boolean.expr_of_bool_expr x) in
+ Z3native.fixedpoint_to_string (z3obj_gnc x) (z3obj_gno x) (List.length queries) (Array.of_list (List.map f queries))
+
+ let get_rules ( x : fixedpoint ) =
+ let v = (AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.fixedpoint_get_rules (z3obj_gnc x) (z3obj_gno x))) in
+ let n = (AST.ASTVector.get_size v) in
+ let f i = Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get v i))) in
+ mk_list f n
+
+ let get_assertions ( x : fixedpoint ) =
+ let v = (AST.ASTVector.ast_vector_of_ptr (z3obj_gc x) (Z3native.fixedpoint_get_assertions (z3obj_gnc x) (z3obj_gno x))) in
+ let n = (AST.ASTVector.get_size v) in
+ let f i = Boolean.bool_expr_of_expr (expr_of_ptr (z3obj_gc x) (z3obj_gno (AST.ASTVector.get v i))) in
+ mk_list f n
+
+ let mk_fixedpoint ( ctx : context ) = create ctx
+end
+
+module Options =
+struct
+
+ let update_param_value ( ctx : context ) ( id : string) ( value : string )=
+ Z3native.update_param_value (context_gno ctx) id value
+
+ let get_param_value ( ctx : context ) ( id : string ) =
+ let (r, v) = (Z3native.get_param_value (context_gno ctx) id) in
+ if not r then
+ None
+ else
+ Some v
+
+ let set_print_mode ( ctx : context ) ( value : ast_print_mode ) =
+ Z3native.set_ast_print_mode (context_gno ctx) (int_of_ast_print_mode value)
+
+ let toggle_warning_messages ( enabled: bool ) =
+ Z3native.toggle_warning_messages enabled
+end
+
+
+module SMT =
+struct
+ let benchmark_to_smtstring ( ctx : context ) ( name : string ) ( logic : string ) ( status : string ) ( attributes : string ) ( assumptions : Boolean.bool_expr list ) ( formula : Boolean.bool_expr ) =
+ Z3native.benchmark_to_smtlib_string (context_gno ctx) name logic status attributes
+ (List.length assumptions) (let f x = ptr_of_expr (Boolean.expr_of_bool_expr x) in (Array.of_list (List.map f assumptions)))
+ (Boolean.gno formula)
+
+ let parse_smtlib_string ( ctx : context ) ( str : string ) ( sort_names : Symbol.symbol list ) ( sorts : sort list ) ( decl_names : Symbol.symbol list ) ( decls : func_decl list ) =
+ let csn = (List.length sort_names) in
+ let cs = (List.length sorts) in
+ let cdn = (List.length decl_names) in
+ let cd = (List.length decls) in
+ if (csn != cs || cdn != cd) then
+ raise (Z3native.Exception "Argument size mismatch")
+ else
+ Z3native.parse_smtlib_string (context_gno ctx) str
+ cs
+ (Symbol.symbol_lton sort_names)
+ (sort_lton sorts)
+ cd
+ (Symbol.symbol_lton decl_names)
+ (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))
+
+ let parse_smtlib_file ( ctx : context ) ( file_name : string ) ( sort_names : Symbol.symbol list ) ( sorts : sort list ) ( decl_names : Symbol.symbol list ) ( decls : func_decl list ) =
+ let csn = (List.length sort_names) in
+ let cs = (List.length sorts) in
+ let cdn = (List.length decl_names) in
+ let cd = (List.length decls) in
+ if (csn != cs || cdn != cd) then
+ raise (Z3native.Exception "Argument size mismatch")
+ else
+ Z3native.parse_smtlib_file (context_gno ctx) file_name
+ cs
+ (Symbol.symbol_lton sort_names)
+ (sort_lton sorts)
+ cd
+ (Symbol.symbol_lton decl_names)
+ (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))
+
+ let get_num_smtlib_formulas ( ctx : context ) = Z3native.get_smtlib_num_formulas (context_gno ctx)
+
+ let get_smtlib_formulas ( ctx : context ) =
+ let n = (get_num_smtlib_formulas ctx ) in
+ let f i = Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.get_smtlib_formula (context_gno ctx) i)) in
+ mk_list f n
+
+ let get_num_smtlib_assumptions ( ctx : context ) = Z3native.get_smtlib_num_assumptions (context_gno ctx)
+
+ let get_smtlib_assumptions ( ctx : context ) =
+ let n = (get_num_smtlib_assumptions ctx ) in
+ let f i = Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.get_smtlib_assumption (context_gno ctx) i)) in
+ mk_list f n
+
+ let get_num_smtlib_decls ( ctx : context ) = Z3native.get_smtlib_num_decls (context_gno ctx)
+
+ let get_smtlib_decls ( ctx : context ) =
+ let n = (get_num_smtlib_decls ctx) in
+ let f i = func_decl_of_ptr ctx (Z3native.get_smtlib_decl (context_gno ctx) i) in
+ mk_list f n
+
+ let get_num_smtlib_sorts ( ctx : context ) = Z3native.get_smtlib_num_sorts (context_gno ctx)
+
+ let get_smtlib_sorts ( ctx : context ) =
+ let n = (get_num_smtlib_sorts ctx) in
+ let f i = (sort_of_ptr ctx (Z3native.get_smtlib_sort (context_gno ctx) i)) in
+ mk_list f n
+
+ let parse_smtlib2_string ( ctx : context ) ( str : string ) ( sort_names : Symbol.symbol list ) ( sorts : sort list ) ( decl_names : Symbol.symbol list ) ( decls : func_decl list ) =
+ let csn = (List.length sort_names) in
+ let cs = (List.length sorts) in
+ let cdn = (List.length decl_names) in
+ let cd = (List.length decls) in
+ if (csn != cs || cdn != cd) then
+ raise (Z3native.Exception "Argument size mismatch")
+ else
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.parse_smtlib2_string (context_gno ctx) str
+ cs
+ (Symbol.symbol_lton sort_names)
+ (sort_lton sorts)
+ cd
+ (Symbol.symbol_lton decl_names)
+ (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))))
+
+ let parse_smtlib2_file ( ctx : context ) ( file_name : string ) ( sort_names : Symbol.symbol list ) ( sorts : sort list ) ( decl_names : Symbol.symbol list ) ( decls : func_decl list ) =
+ let csn = (List.length sort_names) in
+ let cs = (List.length sorts) in
+ let cdn = (List.length decl_names) in
+ let cd = (List.length decls) in
+ if (csn != cs || cdn != cd) then
+ raise (Z3native.Exception "Argument size mismatch")
+ else
+ Boolean.bool_expr_of_expr (expr_of_ptr ctx (Z3native.parse_smtlib2_string (context_gno ctx) file_name
+ cs
+ (Symbol.symbol_lton sort_names)
+ (sort_lton sorts)
+ cd
+ (Symbol.symbol_lton decl_names)
+ (let f x = FuncDecl.gno x in (Array.of_list (List.map f decls)))))
+end
+
+
+let set_global_param ( id : string ) ( value : string ) =
+ (Z3native.global_param_set id value)
+
+let get_global_param ( id : string ) =
+ let (r, v) = (Z3native.global_param_get id) in
+ if not r then
+ None
+ else
+ Some v
+
+let global_param_reset_all =
+ Z3native.global_param_reset_all
diff --git a/src/api/ml/z3_rich.mli b/src/api/ml/z3_rich.mli
new file mode 100644
index 000000000..8b6681e11
--- /dev/null
+++ b/src/api/ml/z3_rich.mli
@@ -0,0 +1,3072 @@
+(**
+ The Z3 ML/Ocaml Interface.
+
+ Copyright (C) 2012 Microsoft Corporation
+ @author CM Wintersteiger (cwinter) 2012-12-17
+
+ NOTE: This is the *rich* version of the interface, using more
+ type information directly in the type system. Coercion functions
+ are provided to tran coerce on type into another where applicable.
+*)
+
+(** Context objects.
+
+ Most interactions with Z3 are interpreted in some context; many users will only
+ require one such object, but power users may require more than one. To start using
+ Z3, do
+
+
+ let ctx = (mk_context []) in
+ (...)
+
+
+ where a list of pairs of strings may be passed to set options on
+ the context, e.g., like so:
+
+
+ let cfg = [("model", "true"); ("...", "...")] in
+ let ctx = (mk_context cfg) in
+ (...)
+
+*)
+type context
+
+(** Create a context object *)
+val mk_context : (string * string) list -> context
+
+
+(** Interaction logging for Z3
+ Note that this is a global, static log and if multiple Context
+ objects are created, it logs the interaction with all of them. *)
+module Log :
+sig
+ (** Open an interaction log file.
+ @return True if opening the log file succeeds, false otherwise. *)
+ (* CMW: "open" seems to be a reserved keyword? *)
+ val open_ : string -> bool
+
+ (** Closes the interaction log. *)
+ val close : unit
+
+ (** Appends a user-provided string to the interaction log. *)
+ val append : string -> unit
+end
+
+(** Version information *)
+module Version :
+sig
+ (** The major version. *)
+ val major : int
+
+ (** The minor version. *)
+ val minor : int
+
+ (** The build version. *)
+ val build : int
+
+ (** The revision. *)
+ val revision : int
+
+ (** A string representation of the version information. *)
+ val to_string : string
+end
+
+(** Symbols are used to name several term and type constructors *)
+module Symbol :
+sig
+ (** Numbered Symbols *)
+ type int_symbol
+
+ (** Named Symbols *)
+ type string_symbol
+
+ (** Symbols *)
+ type symbol = S_Int of int_symbol | S_Str of string_symbol
+
+ (** The kind of the symbol (int or string) *)
+ val kind : symbol -> Z3enums.symbol_kind
+
+ (** Indicates whether the symbol is of Int kind *)
+ val is_int_symbol : symbol -> bool
+
+ (** Indicates whether the symbol is of string kind. *)
+ val is_string_symbol : symbol -> bool
+
+ (** The int value of the symbol. *)
+ val get_int : int_symbol -> int
+
+ (** The string value of the symbol. *)
+ val get_string : string_symbol -> string
+
+ (** A string representation of the symbol. *)
+ val to_string : symbol -> string
+
+ (** Creates a new symbol using an integer.
+
+ Not all integers can be passed to this function.
+ The legal range of unsigned integers is 0 to 2^30-1. *)
+ val mk_int : context -> int -> symbol
+
+ (** Creates a new symbol using a string. *)
+ val mk_string : context -> string -> symbol
+
+ (** Create a list of symbols. *)
+ val mk_ints : context -> int list -> symbol list
+
+ (** Create a list of symbols. *)
+ val mk_strings : context -> string list -> symbol list
+end
+
+(** The abstract syntax tree (AST) module *)
+module AST :
+sig
+ type ast
+
+ (** Vectors of ASTs *)
+ module ASTVector :
+ sig
+ type ast_vector
+
+ (** The size of the vector *)
+ val get_size : ast_vector -> int
+
+ (**
+ Retrieves the i-th object in the vector.
+ @return An AST *)
+ val get : ast_vector -> int -> ast
+
+ (** Sets the i-th object in the vector. *)
+ val set : ast_vector -> int -> ast -> unit
+
+ (** Resize the vector to a new size. *)
+ val resize : ast_vector -> int -> unit
+
+ (** Add an ast to the back of the vector. The size
+ is increased by 1. *)
+ val push : ast_vector -> ast -> unit
+
+ (** Translates all ASTs in the vector to another context.
+ @return A new ASTVector *)
+ val translate : ast_vector -> context -> ast_vector
+
+ (** Retrieves a string representation of the vector. *)
+ val to_string : ast_vector -> string
+ end
+
+ (** Map from AST to AST *)
+ module ASTMap :
+ sig
+ type ast_map
+
+ (** Checks whether the map contains a key.
+ @return True if the key in the map, false otherwise. *)
+ val contains : ast_map -> ast -> bool
+
+ (** Finds the value associated with the key.
+ This function signs an error when the key is not a key in the map. *)
+ val find : ast_map -> ast -> ast
+
+ (** Stores or replaces a new key/value pair in the map. *)
+ val insert : ast_map -> ast -> ast -> unit
+
+ (** Erases the key from the map.*)
+ val erase : ast_map -> ast -> unit
+
+ (** Removes all keys from the map. *)
+ val reset : ast_map -> unit
+
+ (** The size of the map *)
+ val get_size : ast_map -> int
+
+ (** The keys stored in the map. *)
+ val get_keys : ast_map -> ast list
+
+ (** Retrieves a string representation of the map.*)
+ val to_string : ast_map -> string
+ end
+
+ (** The AST's hash code.
+ @return A hash code *)
+ val get_hash_code : ast -> int
+
+ (** A unique identifier for the AST (unique among all ASTs). *)
+ val get_id : ast -> int
+
+ (** The kind of the AST. *)
+ val get_ast_kind : ast -> Z3enums.ast_kind
+
+ (** Indicates whether the AST is an Expr *)
+ val is_expr : ast -> bool
+
+ (** Indicates whether the AST is a bound variable*)
+ val is_var : ast -> bool
+
+ (** Indicates whether the AST is a Quantifier *)
+ val is_quantifier : ast -> bool
+
+ (** Indicates whether the AST is a Sort *)
+ val is_sort : ast -> bool
+
+ (** Indicates whether the AST is a func_decl *)
+ val is_func_decl : ast -> bool
+
+ (** A string representation of the AST. *)
+ val to_string : ast -> string
+
+ (** A string representation of the AST in s-expression notation. *)
+ val to_sexpr : ast -> string
+
+ (** Comparison operator.
+ @return True if the two ast's are from the same context
+ and represent the same sort; false otherwise. *)
+ val ( = ) : ast -> ast -> bool
+
+ (** Object Comparison.
+ @return Negative if the first ast should be sorted before the second, positive if after else zero. *)
+ val compare : ast -> ast -> int
+
+ (** Operator < *)
+ val ( < ) : ast -> ast -> int
+
+ (** Translates (copies) the AST to another context.
+ @return A copy of the AST which is associated with the other context. *)
+ val translate : ast -> context -> ast
+
+ (** Wraps an AST.
+
+ This function is used for transitions between native and
+ managed objects. Note that the native ast that is passed must be a
+ native object obtained from Z3 (e.g., through {!unwrap_ast})
+ and that it must have a correct reference count (see e.g.,
+ Z3native.inc_ref). *)
+ val wrap_ast : context -> Z3native.z3_ast -> ast
+
+ (** Unwraps an AST.
+ This function is used for transitions between native and
+ managed objects. It returns the native pointer to the AST. Note that
+ AST objects are reference counted and unwrapping an AST disables automatic
+ reference counting, i.e., all references to the IntPtr that is returned
+ must be handled externally and through native calls (see e.g.,
+ Z3native.inc_ref).
+ {!wrap_ast} *)
+ val unwrap_ast : ast -> Z3native.ptr
+end
+
+(** The Sort module implements type information for ASTs *)
+module Sort :
+sig
+ (** Sorts *)
+ type sort = Sort of AST.ast
+
+ (** Uninterpreted Sorts *)
+ type uninterpreted_sort = UninterpretedSort of sort
+
+ val ast_of_sort : sort -> AST.ast
+ val sort_of_uninterpreted_sort : uninterpreted_sort -> sort
+ val uninterpreted_sort_of_sort : sort -> uninterpreted_sort
+
+ (** Comparison operator.
+ @return True if the two sorts are from the same context
+ and represent the same sort; false otherwise. *)
+ val ( = ) : sort -> sort -> bool
+
+ (** Returns a unique identifier for the sort. *)
+ val get_id : sort -> int
+
+ (** The kind of the sort. *)
+ val get_sort_kind : sort -> Z3enums.sort_kind
+
+ (** The name of the sort *)
+ val get_name : sort -> Symbol.symbol
+
+ (** A string representation of the sort. *)
+ val to_string : sort -> string
+
+ (** Create a new uninterpreted sort. *)
+ val mk_uninterpreted : context -> Symbol.symbol -> uninterpreted_sort
+
+ (** Create a new uninterpreted sort. *)
+ val mk_uninterpreted_s : context -> string -> uninterpreted_sort
+end
+
+(** Function declarations *)
+module rec FuncDecl :
+sig
+ type func_decl = FuncDecl of AST.ast
+
+ val ast_of_func_decl : FuncDecl.func_decl -> AST.ast
+
+ (** Parameters of Func_Decls *)
+ module Parameter :
+ sig
+ (** Parameters of func_decls *)
+ type parameter =
+ P_Int of int
+ | P_Dbl of float
+ | P_Sym of Symbol.symbol
+ | P_Srt of Sort.sort
+ | P_Ast of AST.ast
+ | P_Fdl of func_decl
+ | P_Rat of string
+
+ (** The kind of the parameter. *)
+ val get_kind : parameter -> Z3enums.parameter_kind
+
+ (** The int value of the parameter.*)
+ val get_int : parameter -> int
+
+ (** The double value of the parameter.*)
+ val get_float : parameter -> float
+
+ (** The Symbol.Symbol value of the parameter.*)
+ val get_symbol : parameter -> Symbol.symbol
+
+ (** The Sort value of the parameter.*)
+ val get_sort : parameter -> Sort.sort
+
+ (** The AST value of the parameter.*)
+ val get_ast : parameter -> AST.ast
+
+ (** The FunctionDeclaration value of the parameter.*)
+ val get_func_decl : parameter -> func_decl
+
+ (** The rational string value of the parameter.*)
+ val get_rational : parameter -> string
+ end
+
+ (** Creates a new function declaration. *)
+ val mk_func_decl : context -> Symbol.symbol -> Sort.sort list -> Sort.sort -> func_decl
+
+ (** Creates a new function declaration. *)
+ val mk_func_decl_s : context -> string -> Sort.sort list -> Sort.sort -> func_decl
+ (** Creates a fresh function declaration with a name prefixed with a prefix string. *)
+
+ val mk_fresh_func_decl : context -> string -> Sort.sort list -> Sort.sort -> func_decl
+
+ (** Creates a new constant function declaration. *)
+ val mk_const_decl : context -> Symbol.symbol -> Sort.sort -> func_decl
+
+ (** Creates a new constant function declaration. *)
+ val mk_const_decl_s : context -> string -> Sort.sort -> func_decl
+
+ (** Creates a fresh constant function declaration with a name prefixed with a prefix string.
+ {!mk_func_decl}
+ {!mk_func_decl} *)
+ val mk_fresh_const_decl : context -> string -> Sort.sort -> func_decl
+
+ (** Comparison operator.
+ @return True if a and b are from the same context and represent the same func_decl; false otherwise. *)
+ val ( = ) : func_decl -> func_decl -> bool
+
+ (** A string representations of the function declaration. *)
+ val to_string : func_decl -> string
+
+ (** Returns a unique identifier for the function declaration. *)
+ val get_id : func_decl -> int
+
+ (** The arity of the function declaration *)
+ val get_arity : func_decl -> int
+
+ (** The size of the domain of the function declaration
+ {!get_arity} *)
+ val get_domain_size : func_decl -> int
+
+ (** The domain of the function declaration *)
+ val get_domain : func_decl -> Sort.sort list
+
+ (** The range of the function declaration *)
+ val get_range : func_decl -> Sort.sort
+
+ (** The kind of the function declaration. *)
+ val get_decl_kind : func_decl -> Z3enums.decl_kind
+
+ (** The name of the function declaration*)
+ val get_name : func_decl -> Symbol.symbol
+
+ (** The number of parameters of the function declaration *)
+ val get_num_parameters : func_decl -> int
+
+ (** The parameters of the function declaration *)
+ val get_parameters : func_decl -> Parameter.parameter list
+
+ (** Create expression that applies function to arguments. *)
+ val apply : func_decl -> Expr.expr list -> Expr.expr
+end
+
+(** Parameter sets (of Solvers, Tactics, ...)
+
+ A Params objects represents a configuration in the form of Symbol.symbol/value pairs. *)
+and Params :
+sig
+ type params
+
+ (** ParamDescrs describe sets of parameters (of Solvers, Tactics, ...) *)
+ module ParamDescrs :
+ sig
+ type param_descrs
+
+ (** Validate a set of parameters. *)
+ val validate : param_descrs -> params -> unit
+
+ (** Retrieve kind of parameter. *)
+ val get_kind : param_descrs -> Symbol.symbol -> Z3enums.param_kind
+
+ (** Retrieve all names of parameters. *)
+ val get_names : param_descrs -> Symbol.symbol list
+
+ (** The size of the ParamDescrs. *)
+ val get_size : param_descrs -> int
+
+ (** Retrieves a string representation of the ParamDescrs. *)
+ val to_string : param_descrs -> string
+ end
+
+ (** Adds a parameter setting. *)
+ val add_bool : params -> Symbol.symbol -> bool -> unit
+
+ (** Adds a parameter setting. *)
+ val add_int : params -> Symbol.symbol -> int -> unit
+
+ (** Adds a parameter setting. *)
+ val add_double : params -> Symbol.symbol -> float -> unit
+
+ (** Adds a parameter setting. *)
+ val add_symbol : params -> Symbol.symbol -> Symbol.symbol -> unit
+
+ (** Adds a parameter setting. *)
+ val add_s_bool : params -> string -> bool -> unit
+
+ (** Adds a parameter setting. *)
+ val add_s_int : params -> string -> int -> unit
+
+ (** Adds a parameter setting. *)
+ val add_s_double : params -> string -> float -> unit
+
+ (** Adds a parameter setting. *)
+ val add_s_symbol : params -> string -> Symbol.symbol -> unit
+
+ (** Creates a new parameter set *)
+ val mk_params : context -> params
+
+ (** A string representation of the parameter set. *)
+ val to_string : params -> string
+end
+
+(** General Expressions (terms) *)
+and Expr :
+sig
+ type expr = Expr of AST.ast
+
+ val ast_of_expr : Expr.expr -> AST.ast
+ val expr_of_ast : AST.ast -> Expr.expr
+
+ (** Returns a simplified version of the expression.
+ {!get_simplify_help} *)
+ val simplify : Expr.expr -> Params.params option -> expr
+
+ (** A string describing all available parameters to Expr.Simplify. *)
+ val get_simplify_help : context -> string
+
+ (** Retrieves parameter descriptions for simplifier. *)
+ val get_simplify_parameter_descrs : context -> Params.ParamDescrs.param_descrs
+
+ (** The function declaration of the function that is applied in this expression. *)
+ val get_func_decl : Expr.expr -> FuncDecl.func_decl
+
+ (** Indicates whether the expression is the true or false expression
+ or something else (L_UNDEF). *)
+ val get_bool_value : Expr.expr -> Z3enums.lbool
+
+ (** The number of arguments of the expression. *)
+ val get_num_args : Expr.expr -> int
+
+ (** The arguments of the expression. *)
+ val get_args : Expr.expr -> Expr.expr list
+
+ (** Update the arguments of the expression using an array of expressions.
+ The number of new arguments should coincide with the current number of arguments. *)
+ val update : Expr.expr -> Expr.expr list -> expr
+
+ (** Substitute every occurrence of from[i] in the expression with to[i], for i smaller than num_exprs.
+
+ The result is the new expression. The arrays from and to must have size num_exprs.
+ For every i smaller than num_exprs, we must have that
+ sort of from[i] must be equal to sort of to[i]. *)
+ val substitute : Expr.expr -> Expr.expr list -> Expr.expr list -> expr
+
+ (** Substitute every occurrence of from in the expression with to.
+ {!substitute} *)
+ val substitute_one : Expr.expr -> Expr.expr -> Expr.expr -> expr
+
+ (** Substitute the free variables in the expression with the expressions in the expr array
+
+ For every i smaller than num_exprs, the variable with de-Bruijn index i is replaced with term to[i]. *)
+ val substitute_vars : Expr.expr -> Expr.expr list -> expr
+
+ (** Translates (copies) the term to another context.
+ @return A copy of the term which is associated with the other context *)
+ val translate : Expr.expr -> context -> expr
+
+ (** Returns a string representation of the expression. *)
+ val to_string : Expr.expr -> string
+
+ (** Indicates whether the term is a numeral *)
+ val is_numeral : Expr.expr -> bool
+
+ (** Indicates whether the term is well-sorted.
+ @return True if the term is well-sorted, false otherwise. *)
+ val is_well_sorted : Expr.expr -> bool
+
+ (** The Sort of the term. *)
+ val get_sort : Expr.expr -> Sort.sort
+
+ (** Indicates whether the term has Boolean sort. *)
+ val is_bool : Expr.expr -> bool
+
+ (** Indicates whether the term represents a constant. *)
+ val is_const : Expr.expr -> bool
+
+ (** Indicates whether the term is the constant true. *)
+ val is_true : Expr.expr -> bool
+
+ (** Indicates whether the term is the constant false. *)
+ val is_false : Expr.expr -> bool
+
+ (** Indicates whether the term is an equality predicate. *)
+ val is_eq : Expr.expr -> bool
+
+ (** Indicates whether the term is an n-ary distinct predicate (every argument is mutually distinct). *)
+ val is_distinct : Expr.expr -> bool
+
+ (** Indicates whether the term is a ternary if-then-else term *)
+ val is_ite : Expr.expr -> bool
+
+ (** Indicates whether the term is an n-ary conjunction *)
+ val is_and : Expr.expr -> bool
+
+ (** Indicates whether the term is an n-ary disjunction *)
+ val is_or : Expr.expr -> bool
+
+ (** Indicates whether the term is an if-and-only-if (Boolean equivalence, binary) *)
+ val is_iff : Expr.expr -> bool
+
+ (** Indicates whether the term is an exclusive or *)
+ val is_xor : Expr.expr -> bool
+
+ (** Indicates whether the term is a negation *)
+ val is_not : Expr.expr -> bool
+
+ (** Indicates whether the term is an implication *)
+ val is_implies : Expr.expr -> bool
+
+ (** Indicates whether the term is a label (used by the Boogie Verification condition generator).
+ The label has two parameters, a string and a Boolean polarity. It takes one argument, a formula. *)
+ val is_label : Expr.expr -> bool
+
+ (** Indicates whether the term is a label literal (used by the Boogie Verification condition generator).
+ A label literal has a set of string parameters. It takes no arguments.
+ let is_label_lit ( x : expr ) = (FuncDecl.get_decl_kind (get_func_decl x) == OP_LABEL_LIT) *)
+ val is_label_lit : Expr.expr -> bool
+
+ (** Indicates whether the term is a binary equivalence modulo namings.
+ This binary predicate is used in proof terms.
+ It captures equisatisfiability and equivalence modulo renamings. *)
+ val is_oeq : Expr.expr -> bool
+
+ (** Creates a new constant. *)
+ val mk_const : context -> Symbol.symbol -> Sort.sort -> expr
+
+ (** Creates a new constant. *)
+ val mk_const_s : context -> string -> Sort.sort -> expr
+
+ (** Creates a constant from the func_decl. *)
+ val mk_const_f : context -> FuncDecl.func_decl -> expr
+
+ (** Creates a fresh constant with a name prefixed with a string. *)
+ val mk_fresh_const : context -> string -> Sort.sort -> expr
+
+ (** Create a new function application. *)
+ val mk_app : context -> FuncDecl.func_decl -> Expr.expr list -> expr
+
+ (** Create a numeral of a given sort.
+ @return A Term with the goven value and sort *)
+ val mk_numeral_string : context -> string -> Sort.sort -> expr
+
+ (** Create a numeral of a given sort. This function can be use to create numerals that fit in a machine integer.
+ It is slightly faster than MakeNumeral since it is not necessary to parse a string.
+ @return A Term with the given value and sort *)
+ val mk_numeral_int : context -> int -> Sort.sort -> expr
+end
+
+(** Boolean expressions *)
+module Boolean :
+sig
+ type bool_sort = BoolSort of Sort.sort
+ type bool_expr = BoolExpr of Expr.expr
+
+ val expr_of_bool_expr : bool_expr -> Expr.expr
+ val sort_of_bool_sort : bool_sort -> Sort.sort
+ val bool_sort_of_sort : Sort.sort -> bool_sort
+ val bool_expr_of_expr : Expr.expr -> bool_expr
+
+ (** Create a Boolean sort *)
+ val mk_sort : context -> bool_sort
+
+ (** Create a Boolean constant. *)
+ val mk_const : context -> Symbol.symbol -> bool_expr
+
+ (** Create a Boolean constant. *)
+ val mk_const_s : context -> string -> bool_expr
+
+ (** The true Term. *)
+ val mk_true : context -> bool_expr
+
+ (** The false Term. *)
+ val mk_false : context -> bool_expr
+
+ (** Creates a Boolean value. *)
+ val mk_val : context -> bool -> bool_expr
+
+ (** Creates the equality between two expr's. *)
+ val mk_eq : context -> Expr.expr -> Expr.expr -> bool_expr
+
+ (** Creates a distinct term. *)
+ val mk_distinct : context -> Expr.expr list -> bool_expr
+
+ (** Mk an expression representing not(a). *)
+ val mk_not : context -> bool_expr -> bool_expr
+
+ (** Create an expression representing an if-then-else: ite(t1, t2, t3). *)
+ val mk_ite : context -> bool_expr -> bool_expr -> bool_expr -> bool_expr
+
+ (** Create an expression representing t1 iff t2. *)
+ val mk_iff : context -> bool_expr -> bool_expr -> bool_expr
+
+ (** Create an expression representing t1 -> t2. *)
+ val mk_implies : context -> bool_expr -> bool_expr -> bool_expr
+
+ (** Create an expression representing t1 xor t2. *)
+ val mk_xor : context -> bool_expr -> bool_expr -> bool_expr
+
+ (** Create an expression representing the AND of args *)
+ val mk_and : context -> bool_expr list -> bool_expr
+
+ (** Create an expression representing the OR of args *)
+ val mk_or : context -> bool_expr list -> bool_expr
+end
+
+(** Quantifier expressions *)
+module Quantifier :
+sig
+ type quantifier = Quantifier of Expr.expr
+
+ val expr_of_quantifier : quantifier -> Expr.expr
+ val quantifier_of_expr : Expr.expr -> quantifier
+
+ (** Quantifier patterns
+
+ Patterns comprise a list of terms. The list should be
+ non-empty. If the list comprises of more than one term, it is
+ also called a multi-pattern. *)
+ module Pattern :
+ sig
+ type pattern = Pattern of AST.ast
+
+ val ast_of_pattern : pattern -> AST.ast
+ val pattern_of_ast : AST.ast -> pattern
+
+ (** The number of terms in the pattern. *)
+ val get_num_terms : pattern -> int
+
+ (** The terms in the pattern. *)
+ val get_terms : pattern -> Expr.expr list
+
+ (** A string representation of the pattern. *)
+ val to_string : pattern -> string
+ end
+
+
+ (** The de-Burijn index of a bound variable.
+
+ Bound variables are indexed by de-Bruijn indices. It is perhaps easiest to explain
+ the meaning of de-Bruijn indices by indicating the compilation process from
+ non-de-Bruijn formulas to de-Bruijn format.
+
+ abs(forall (x1) phi) = forall (x1) abs1(phi, x1, 0)
+ abs(forall (x1, x2) phi) = abs(forall (x1) abs(forall (x2) phi))
+ abs1(x, x, n) = b_n
+ abs1(y, x, n) = y
+ abs1(f(t1,...,tn), x, n) = f(abs1(t1,x,n), ..., abs1(tn,x,n))
+ abs1(forall (x1) phi, x, n) = forall (x1) (abs1(phi, x, n+1))
+
+ The last line is significant: the index of a bound variable is different depending
+ on the scope in which it appears. The deeper ( x : expr ) appears, the higher is its
+ index. *)
+ val get_index : Expr.expr -> int
+
+ (** Indicates whether the quantifier is universal. *)
+ val is_universal : quantifier -> bool
+
+ (** Indicates whether the quantifier is existential. *)
+ val is_existential : quantifier -> bool
+
+ (** The weight of the quantifier. *)
+ val get_weight : quantifier -> int
+
+ (** The number of patterns. *)
+ val get_num_patterns : quantifier -> int
+
+ (** The patterns. *)
+ val get_patterns : quantifier -> Pattern.pattern list
+
+ (** The number of no-patterns. *)
+ val get_num_no_patterns : quantifier -> int
+
+ (** The no-patterns. *)
+ val get_no_patterns : quantifier -> Pattern.pattern list
+
+ (** The number of bound variables. *)
+ val get_num_bound : quantifier -> int
+
+ (** The symbols for the bound variables. *)
+ val get_bound_variable_names : quantifier -> Symbol.symbol list
+
+ (** The sorts of the bound variables. *)
+ val get_bound_variable_sorts : quantifier -> Sort.sort list
+
+ (** The body of the quantifier. *)
+ val get_body : quantifier -> Boolean.bool_expr
+
+ (** Creates a new bound variable. *)
+ val mk_bound : context -> int -> Sort.sort -> Expr.expr
+
+ (** Create a quantifier pattern. *)
+ val mk_pattern : context -> Expr.expr list -> Pattern.pattern
+
+ (** Create a universal Quantifier. *)
+ val mk_forall : context -> Sort.sort list -> Symbol.symbol list -> Expr.expr -> int option -> Pattern.pattern list -> Expr.expr list -> Symbol.symbol option -> Symbol.symbol option -> quantifier
+
+ (** Create a universal Quantifier. *)
+ val mk_forall_const : context -> Expr.expr list -> Expr.expr -> int option -> Pattern.pattern list -> Expr.expr list -> Symbol.symbol option -> Symbol.symbol option -> quantifier
+
+ (** Create an existential Quantifier. *)
+ val mk_exists : context -> Sort.sort list -> Symbol.symbol list -> Expr.expr -> int option -> Pattern.pattern list -> Expr.expr list -> Symbol.symbol option -> Symbol.symbol option -> quantifier
+
+ (** Create an existential Quantifier. *)
+ val mk_exists_const : context -> Expr.expr list -> Expr.expr -> int option -> Pattern.pattern list -> Expr.expr list -> Symbol.symbol option -> Symbol.symbol option -> quantifier
+
+ (** Create a Quantifier. *)
+ val mk_quantifier : context -> Sort.sort list -> Symbol.symbol list -> Expr.expr -> int option -> Pattern.pattern list -> Expr.expr list -> Symbol.symbol option -> Symbol.symbol option -> quantifier
+
+ (** Create a Quantifier. *)
+ val mk_quantifier : context -> bool -> Expr.expr list -> Expr.expr -> int option -> Pattern.pattern list -> Expr.expr list -> Symbol.symbol option -> Symbol.symbol option -> quantifier
+end
+
+(** Functions to manipulate Array expressions *)
+module Array_ :
+sig
+ type array_sort = ArraySort of Sort.sort
+ type array_expr = ArrayExpr of Expr.expr
+
+ val sort_of_array_sort : array_sort -> Sort.sort
+ val array_sort_of_sort : Sort.sort -> array_sort
+ val expr_of_array_expr : array_expr -> Expr.expr
+
+ val array_expr_of_expr : Expr.expr -> array_expr
+
+ (** Create a new array sort. *)
+ val mk_sort : context -> Sort.sort -> Sort.sort -> array_sort
+
+ (** Indicates whether the term is an array store.
+ It satisfies select(store(a,i,v),j) = if i = j then v else select(a,j).
+ Array store takes at least 3 arguments. *)
+ val is_store : Expr.expr -> bool
+
+ (** Indicates whether the term is an array select. *)
+ val is_select : Expr.expr -> bool
+
+ (** Indicates whether the term is a constant array.
+ For example, select(const(v),i) = v holds for every v and i. The function is unary. *)
+ val is_constant_array : Expr.expr -> bool
+
+ (** Indicates whether the term is a default array.
+ For example default(const(v)) = v. The function is unary. *)
+ val is_default_array : Expr.expr -> bool
+
+ (** Indicates whether the term is an array map.
+ It satisfies map[f](a1,..,a_n)[i] = f(a1[i],...,a_n[i]) for every i. *)
+ val is_array_map : Expr.expr -> bool
+
+ (** Indicates whether the term is an as-array term.
+ An as-array term is n array value that behaves as the function graph of the
+ function passed as parameter. *)
+ val is_as_array : Expr.expr -> bool
+
+ (** Indicates whether the term is of an array sort. *)
+ val is_array : Expr.expr -> bool
+
+ (** The domain of the array sort. *)
+ val get_domain : array_sort -> Sort.sort
+
+ (** The range of the array sort. *)
+ val get_range : array_sort -> Sort.sort
+
+ (** Create an array constant. *)
+ val mk_const : context -> Symbol.symbol -> Sort.sort -> Sort.sort -> array_expr
+
+ (** Create an array constant. *)
+ val mk_const_s : context -> string -> Sort.sort -> Sort.sort -> array_expr
+
+ (** Array read.
+
+ The argument a is the array and i is the index
+ of the array that gets read.
+
+ The node a must have an array sort [domain -> range],
+ and i must have the sort domain.
+ The sort of the result is range.
+ {!Array_.mk_sort}
+ {!mk_store} *)
+ val mk_select : context -> array_expr -> Expr.expr -> array_expr
+
+ (** Array update.
+
+ The node a must have an array sort [domain -> range],
+ i must have sort domain,
+ v must have sort range. The sort of the result is [domain -> range].
+ The semantics of this function is given by the theory of arrays described in the SMT-LIB
+ standard. See http://smtlib.org for more details.
+ The result of this function is an array that is equal to a
+ (with respect to select)
+ on all indices except for i, where it maps to v
+ (and the select of a with
+ respect to i may be a different value).
+ {!Array_.mk_sort}
+ {!mk_select} *)
+ val mk_store : context -> array_expr -> Expr.expr -> Expr.expr -> array_expr
+
+ (** Create a constant array.
+
+ The resulting term is an array, such that a selecton an arbitrary index
+ produces the value v.
+ {!Array_.mk_sort}
+ {!mk_select} *)
+ val mk_const_array : context -> Sort.sort -> Expr.expr -> array_expr
+
+ (** Maps f on the argument arrays.
+
+ Eeach element of args must be of an array sort [domain_i -> range_i].
+ The function declaration f must have type range_1 .. range_n -> range.
+ v must have sort range. The sort of the result is [domain_i -> range].
+ {!Array_.mk_sort}
+ {!mk_select}
+ {!mk_store} *)
+ val mk_map : context -> FuncDecl.func_decl -> array_expr list -> array_expr
+
+ (** Access the array default value.
+
+ Produces the default range value, for arrays that can be represented as
+ finite maps with a default range value. *)
+ val mk_term_array : context -> array_expr -> array_expr
+end
+
+(** Functions to manipulate Set expressions *)
+module Set :
+sig
+ type set_sort = SetSort of Sort.sort
+
+ val sort_of_set_sort : set_sort -> Sort.sort
+
+ (** Create a set type. *)
+ val mk_sort : context -> Sort.sort -> set_sort
+
+ (** Indicates whether the term is set union *)
+ val is_union : Expr.expr -> bool
+
+ (** Indicates whether the term is set intersection *)
+ val is_intersect : Expr.expr -> bool
+
+ (** Indicates whether the term is set difference *)
+ val is_difference : Expr.expr -> bool
+
+ (** Indicates whether the term is set complement *)
+ val is_complement : Expr.expr -> bool
+
+ (** Indicates whether the term is set subset *)
+ val is_subset : Expr.expr -> bool
+
+ (** Create an empty set. *)
+ val mk_empty : context -> Sort.sort -> Expr.expr
+
+ (** Create the full set. *)
+ val mk_full : context -> Sort.sort -> Expr.expr
+
+ (** Add an element to the set. *)
+ val mk_set_add : context -> Expr.expr -> Expr.expr -> Expr.expr
+
+ (** Remove an element from a set. *)
+ val mk_del : context -> Expr.expr -> Expr.expr -> Expr.expr
+
+ (** Take the union of a list of sets. *)
+ val mk_union : context -> Expr.expr list -> Expr.expr
+
+ (** Take the intersection of a list of sets. *)
+ val mk_intersection : context -> Expr.expr list -> Expr.expr
+
+ (** Take the difference between two sets. *)
+ val mk_difference : context -> Expr.expr -> Expr.expr -> Expr.expr
+
+ (** Take the complement of a set. *)
+ val mk_complement : context -> Expr.expr -> Expr.expr
+
+ (** Check for set membership. *)
+ val mk_membership : context -> Expr.expr -> Expr.expr -> Expr.expr
+
+ (** Check for subsetness of sets. *)
+ val mk_subset : context -> Expr.expr -> Expr.expr -> Expr.expr
+end
+
+(** Functions to manipulate Finite Domain expressions *)
+module FiniteDomain :
+sig
+ type finite_domain_sort = FiniteDomainSort of Sort.sort
+
+ val sort_of_finite_domain_sort : finite_domain_sort -> Sort.sort
+ val finite_domain_sort_of_sort : Sort.sort -> finite_domain_sort
+
+ (** Create a new finite domain sort. *)
+ val mk_sort : context -> Symbol.symbol -> int -> finite_domain_sort
+
+ (** Create a new finite domain sort. *)
+ val mk_sort_s : context -> string -> int -> finite_domain_sort
+
+ (** Indicates whether the term is of an array sort. *)
+ val is_finite_domain : Expr.expr -> bool
+
+ (** Indicates whether the term is a less than predicate over a finite domain. *)
+ val is_lt : Expr.expr -> bool
+
+ (** The size of the finite domain sort. *)
+ val get_size : finite_domain_sort -> int
+end
+
+
+(** Functions to manipulate Relation expressions *)
+module Relation :
+sig
+ type relation_sort = RelationSort of Sort.sort
+
+ val sort_of_relation_sort : relation_sort -> Sort.sort
+ val relation_sort_of_sort : Sort.sort -> relation_sort
+
+ (** Indicates whether the term is of a relation sort. *)
+ val is_relation : Expr.expr -> bool
+
+ (** Indicates whether the term is an relation store
+
+ Insert a record into a relation.
+ The function takes n+1 arguments, where the first argument is the relation and the remaining n elements
+ correspond to the n columns of the relation. *)
+ val is_store : Expr.expr -> bool
+
+ (** Indicates whether the term is an empty relation *)
+ val is_empty : Expr.expr -> bool
+
+ (** Indicates whether the term is a test for the emptiness of a relation *)
+ val is_is_empty : Expr.expr -> bool
+
+ (** Indicates whether the term is a relational join *)
+ val is_join : Expr.expr -> bool
+
+ (** Indicates whether the term is the union or convex hull of two relations.
+ The function takes two arguments. *)
+ val is_union : Expr.expr -> bool
+
+ (** Indicates whether the term is the widening of two relations
+ The function takes two arguments. *)
+ val is_widen : Expr.expr -> bool
+
+ (** Indicates whether the term is a projection of columns (provided as numbers in the parameters).
+ The function takes one argument. *)
+ val is_project : Expr.expr -> bool
+
+ (** Indicates whether the term is a relation filter
+
+ Filter (restrict) a relation with respect to a predicate.
+ The first argument is a relation.
+ The second argument is a predicate with free de-Brujin indices
+ corresponding to the columns of the relation.
+ So the first column in the relation has index 0. *)
+ val is_filter : Expr.expr -> bool
+
+ (** Indicates whether the term is an intersection of a relation with the negation of another.
+
+ Intersect the first relation with respect to negation
+ of the second relation (the function takes two arguments).
+ Logically, the specification can be described by a function
+
+ target = filter_by_negation(pos, neg, columns)
+
+ where columns are pairs c1, d1, .., cN, dN of columns from pos and neg, such that
+ target are elements in ( x : expr ) in pos, such that there is no y in neg that agrees with
+ ( x : expr ) on the columns c1, d1, .., cN, dN. *)
+ val is_negation_filter : Expr.expr -> bool
+
+ (** Indicates whether the term is the renaming of a column in a relation
+
+ The function takes one argument.
+ The parameters contain the renaming as a cycle. *)
+ val is_rename : Expr.expr -> bool
+
+ (** Indicates whether the term is the complement of a relation *)
+ val is_complement : Expr.expr -> bool
+
+ (** Indicates whether the term is a relational select
+
+ Check if a record is an element of the relation.
+ The function takes n+1 arguments, where the first argument is a relation,
+ and the remaining n arguments correspond to a record. *)
+ val is_select : Expr.expr -> bool
+
+ (** Indicates whether the term is a relational clone (copy)
+
+ Create a fresh copy (clone) of a relation.
+ The function is logically the identity, but
+ in the context of a register machine allows
+ for terms of kind {!is_union}
+ to perform destructive updates to the first argument. *)
+ val is_clone : Expr.expr -> bool
+
+ (** The arity of the relation sort. *)
+ val get_arity : relation_sort -> int
+
+ (** The sorts of the columns of the relation sort. *)
+ val get_column_sorts : relation_sort -> relation_sort list
+end
+
+(** Functions to manipulate Datatype expressions *)
+module Datatype :
+sig
+ type datatype_sort = DatatypeSort of Sort.sort
+ type datatype_expr = DatatypeExpr of Expr.expr
+
+ val sort_of_datatype_sort : datatype_sort -> Sort.sort
+ val datatype_sort_of_sort : Sort.sort -> datatype_sort
+ val expr_of_datatype_expr : datatype_expr -> Expr.expr
+ val datatype_expr_of_expr : Expr.expr -> datatype_expr
+
+ (** Datatype Constructors *)
+ module Constructor :
+ sig
+ type constructor
+
+ (** The number of fields of the constructor. *)
+ val get_num_fields : constructor -> int
+
+ (** The function declaration of the constructor. *)
+ val get_constructor_decl : constructor -> FuncDecl.func_decl
+
+ (** The function declaration of the tester. *)
+ val get_tester_decl : constructor -> FuncDecl.func_decl
+
+ (** The function declarations of the accessors *)
+ val get_accessor_decls : constructor -> FuncDecl.func_decl list
+ end
+
+ (** Create a datatype constructor.
+ if the corresponding sort reference is 0, then the value in sort_refs should be an index
+ referring to one of the recursive datatypes that is declared. *)
+ val mk_constructor : context -> Symbol.symbol -> Symbol.symbol -> Symbol.symbol list -> Sort.sort list -> int list -> Constructor.constructor
+
+ (** Create a datatype constructor.
+ if the corresponding sort reference is 0, then the value in sort_refs should be an index
+ referring to one of the recursive datatypes that is declared. *)
+ val mk_constructor_s : context -> string -> Symbol.symbol -> Symbol.symbol list -> Sort.sort list -> int list -> Constructor.constructor
+
+ (** Create a new datatype sort. *)
+ val mk_sort : context -> Symbol.symbol -> Constructor.constructor list -> datatype_sort
+
+ (** Create a new datatype sort. *)
+ val mk_sort_s : context -> string -> Constructor.constructor list -> datatype_sort
+
+ (** Create mutually recursive datatypes. *)
+ val mk_sorts : context -> Symbol.symbol list -> Constructor.constructor list list -> datatype_sort list
+
+ (** Create mutually recursive data-types. *)
+ val mk_sorts_s : context -> string list -> Constructor.constructor list list -> datatype_sort list
+
+
+ (** The number of constructors of the datatype sort. *)
+ val get_num_constructors : datatype_sort -> int
+
+ (** The constructors. *)
+ val get_constructors : datatype_sort -> FuncDecl.func_decl list
+
+ (** The recognizers. *)
+ val get_recognizers : datatype_sort -> FuncDecl.func_decl list
+
+ (** The constructor accessors. *)
+ val get_accessors : datatype_sort -> FuncDecl.func_decl list list
+end
+
+(** Functions to manipulate Enumeration expressions *)
+module Enumeration :
+sig
+ type enum_sort = EnumSort of Sort.sort
+
+ val sort_of_enum_sort : enum_sort -> Sort.sort
+
+ (** Create a new enumeration sort. *)
+ val mk_sort : context -> Symbol.symbol -> Symbol.symbol list -> enum_sort
+
+ (** Create a new enumeration sort. *)
+ val mk_sort_s : context -> string -> string list -> enum_sort
+
+ (** The function declarations of the constants in the enumeration. *)
+ val get_const_decls : enum_sort -> FuncDecl.func_decl list
+
+ (** The test predicates for the constants in the enumeration. *)
+ val get_tester_decls : enum_sort -> FuncDecl.func_decl list
+end
+
+(** Functions to manipulate List expressions *)
+module List_ :
+sig
+ type list_sort = ListSort of Sort.sort
+
+ val sort_of_list_sort : list_sort -> Sort.sort
+
+ (** Create a new list sort. *)
+ val mk_sort : context -> Symbol.symbol -> Sort.sort -> list_sort
+
+ (** Create a new list sort. *)
+ val mk_list_s : context -> string -> Sort.sort -> list_sort
+
+ (** The declaration of the nil function of this list sort. *)
+ val get_nil_decl : list_sort -> FuncDecl.func_decl
+
+ (** The declaration of the isNil function of this list sort. *)
+ val get_is_nil_decl : list_sort -> FuncDecl.func_decl
+
+ (** The declaration of the cons function of this list sort. *)
+ val get_cons_decl : list_sort -> FuncDecl.func_decl
+
+ (** The declaration of the isCons function of this list sort. *)
+ val get_is_cons_decl : list_sort -> FuncDecl.func_decl
+
+ (** The declaration of the head function of this list sort. *)
+ val get_head_decl : list_sort -> FuncDecl.func_decl
+
+ (** The declaration of the tail function of this list sort. *)
+ val get_tail_decl : list_sort -> FuncDecl.func_decl
+
+ (** The empty list. *)
+ val nil : list_sort -> Expr.expr
+end
+
+(** Functions to manipulate Tuple expressions *)
+module Tuple :
+sig
+ type tuple_sort = TupleSort of Sort.sort
+
+ val sort_of_tuple_sort : tuple_sort -> Sort.sort
+
+ (** Create a new tuple sort. *)
+ val mk_sort : context -> Symbol.symbol -> Symbol.symbol list -> Sort.sort list -> tuple_sort
+
+ (** The constructor function of the tuple. *)
+ val get_mk_decl : tuple_sort -> FuncDecl.func_decl
+
+ (** The number of fields in the tuple. *)
+ val get_num_fields : tuple_sort -> int
+
+ (** The field declarations. *)
+ val get_field_decls : tuple_sort -> FuncDecl.func_decl list
+end
+
+(** Functions to manipulate arithmetic expressions *)
+module rec Arithmetic :
+sig
+ type arith_sort = ArithSort of Sort.sort
+ type arith_expr = ArithExpr of Expr.expr
+
+ val sort_of_arith_sort : Arithmetic.arith_sort -> Sort.sort
+ val arith_sort_of_sort : Sort.sort -> Arithmetic.arith_sort
+ val expr_of_arith_expr : Arithmetic.arith_expr -> Expr.expr
+ val arith_expr_of_expr : Expr.expr -> Arithmetic.arith_expr
+
+ (** Integer Arithmetic *)
+ module rec Integer :
+ sig
+ type int_sort = IntSort of arith_sort
+ type int_expr = IntExpr of arith_expr
+ type int_num = IntNum of int_expr
+
+ val arith_sort_of_int_sort : Arithmetic.Integer.int_sort -> Arithmetic.arith_sort
+ val int_sort_of_arith_sort : Arithmetic.arith_sort -> Arithmetic.Integer.int_sort
+ val arith_expr_of_int_expr : Arithmetic.Integer.int_expr -> Arithmetic.arith_expr
+ val int_expr_of_int_num : Arithmetic.Integer.int_num -> Arithmetic.Integer.int_expr
+ val int_expr_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.Integer.int_expr
+ val int_num_of_int_expr : Arithmetic.Integer.int_expr -> Arithmetic.Integer.int_num
+
+ (** Create a new integer sort. *)
+ val mk_sort : context -> int_sort
+
+ (** Retrieve the int value. *)
+ val get_int : int_num -> int
+
+ (** Returns a string representation of the numeral. *)
+ val to_string : int_num -> string
+
+ (** Creates an integer constant. *)
+ val mk_int_const : context -> Symbol.symbol -> int_expr
+
+ (** Creates an integer constant. *)
+ val mk_int_const_s : context -> string -> int_expr
+
+ (** Create an expression representing t1 mod t2.
+ The arguments must have int type. *)
+ val mk_mod : context -> int_expr -> int_expr -> int_expr
+
+ (** Create an expression representing t1 rem t2.
+ The arguments must have int type. *)
+ val mk_rem : context -> int_expr -> int_expr -> int_expr
+
+ (** Create an integer numeral. *)
+ val mk_int_numeral_s : context -> string -> int_num
+
+ (** Create an integer numeral.
+ @return A Term with the given value and sort Integer *)
+ val mk_int_numeral_i : context -> int -> int_num
+
+ (** Coerce an integer to a real.
+
+
+ There is also a converse operation exposed. It follows the semantics prescribed by the SMT-LIB standard.
+
+ You can take the floor of a real by creating an auxiliary integer Term k and
+ and asserting MakeInt2Real(k) <= t1 < MkInt2Real(k)+1.
+ The argument must be of integer sort. *)
+ val mk_int2real : context -> int_expr -> Real.real_expr
+
+ (** Create an n-bit bit-vector from an integer argument.
+
+
+ NB. This function is essentially treated as uninterpreted.
+ So you cannot expect Z3 to precisely reflect the semantics of this function
+ when solving constraints with this function.
+
+ The argument must be of integer sort. *)
+ val mk_int2bv : context -> int -> int_expr -> BitVector.bitvec_expr
+ end
+
+ (** Real Arithmetic *)
+ and Real :
+ sig
+ type real_sort = RealSort of arith_sort
+ type real_expr = RealExpr of arith_expr
+ type rat_num = RatNum of real_expr
+
+ val arith_sort_of_real_sort : Arithmetic.Real.real_sort -> Arithmetic.arith_sort
+ val real_sort_of_arith_sort : Arithmetic.arith_sort -> Arithmetic.Real.real_sort
+ val arith_expr_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.arith_expr
+ val real_expr_of_rat_num : Arithmetic.Real.rat_num -> Arithmetic.Real.real_expr
+ val real_expr_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.Real.real_expr
+ val rat_num_of_real_expr : Arithmetic.Real.real_expr -> Arithmetic.Real.rat_num
+
+ (** Create a real sort. *)
+ val mk_sort : context -> real_sort
+
+ (** The numerator of a rational numeral. *)
+ val get_numerator : rat_num -> Integer.int_num
+
+ (** The denominator of a rational numeral. *)
+ val get_denominator : rat_num -> Integer.int_num
+
+ (** Returns a string representation in decimal notation.
+ The result has at most as many decimal places as indicated by the int argument.*)
+ val to_decimal_string : rat_num -> int -> string
+
+ (** Returns a string representation of the numeral. *)
+ val to_string : rat_num -> string
+
+ (** Creates a real constant. *)
+ val mk_real_const : context -> Symbol.symbol -> real_expr
+
+ (** Creates a real constant. *)
+ val mk_real_const_s : context -> string -> real_expr
+
+ (** Create a real numeral from a fraction.
+ @return A Term with rational value and sort Real
+ {!mk_numeral_s} *)
+ val mk_numeral_nd : context -> int -> int -> rat_num
+
+ (** Create a real numeral.
+ @return A Term with the given value and sort Real *)
+ val mk_numeral_s : context -> string -> rat_num
+
+ (** Create a real numeral.
+ @return A Term with the given value and sort Real *)
+ val mk_numeral_i : context -> int -> rat_num
+
+ (** Creates an expression that checks whether a real number is an integer. *)
+ val mk_is_integer : context -> real_expr -> Boolean.bool_expr
+
+ (** Coerce a real to an integer.
+
+ The semantics of this function follows the SMT-LIB standard for the function to_int.
+ The argument must be of real sort. *)
+ val mk_real2int : context -> real_expr -> Integer.int_expr
+ end
+
+ (** Algebraic Numbers *)
+ and AlgebraicNumber :
+ sig
+ type algebraic_num = AlgebraicNum of arith_expr
+
+ val arith_expr_of_algebraic_num : Arithmetic.AlgebraicNumber.algebraic_num -> Arithmetic.arith_expr
+ val algebraic_num_of_arith_expr : Arithmetic.arith_expr -> Arithmetic.AlgebraicNumber.algebraic_num
+
+ (** Return a upper bound for a given real algebraic number.
+ The interval isolating the number is smaller than 1/10^precision.
+ {!is_algebraic_number}
+ @return A numeral Expr of sort Real *)
+ val to_upper : algebraic_num -> int -> Real.rat_num
+
+ (** Return a lower bound for the given real algebraic number.
+ The interval isolating the number is smaller than 1/10^precision.
+ {!is_algebraic_number}
+ @return A numeral Expr of sort Real *)
+ val to_lower : algebraic_num -> int -> Real.rat_num
+
+ (** Returns a string representation in decimal notation.
+ The result has at most as many decimal places as the int argument provided.*)
+ val to_decimal_string : algebraic_num -> int -> string
+
+ (** Returns a string representation of the numeral. *)
+ val to_string : algebraic_num -> string
+ end
+
+ (** Indicates whether the term is of integer sort. *)
+ val is_int : Expr.expr -> bool
+
+ (** Indicates whether the term is an arithmetic numeral. *)
+ val is_arithmetic_numeral : Expr.expr -> bool
+
+ (** Indicates whether the term is a less-than-or-equal *)
+ val is_le : Expr.expr -> bool
+
+ (** Indicates whether the term is a greater-than-or-equal *)
+ val is_ge : Expr.expr -> bool
+
+ (** Indicates whether the term is a less-than *)
+ val is_lt : Expr.expr -> bool
+
+ (** Indicates whether the term is a greater-than *)
+ val is_gt : Expr.expr -> bool
+
+ (** Indicates whether the term is addition (binary) *)
+ val is_add : Expr.expr -> bool
+
+ (** Indicates whether the term is subtraction (binary) *)
+ val is_sub : Expr.expr -> bool
+
+ (** Indicates whether the term is a unary minus *)
+ val is_uminus : Expr.expr -> bool
+
+ (** Indicates whether the term is multiplication (binary) *)
+ val is_mul : Expr.expr -> bool
+
+ (** Indicates whether the term is division (binary) *)
+ val is_div : Expr.expr -> bool
+
+ (** Indicates whether the term is integer division (binary) *)
+ val is_idiv : Expr.expr -> bool
+
+ (** Indicates whether the term is remainder (binary) *)
+ val is_remainder : Expr.expr -> bool
+
+ (** Indicates whether the term is modulus (binary) *)
+ val is_modulus : Expr.expr -> bool
+
+ (** Indicates whether the term is a coercion of integer to real (unary) *)
+ val is_inttoreal : Expr.expr -> bool
+
+ (** Indicates whether the term is a coercion of real to integer (unary) *)
+ val is_real_to_int : Expr.expr -> bool
+
+ (** Indicates whether the term is a check that tests whether a real is integral (unary) *)
+ val is_real_is_int : Expr.expr -> bool
+
+ (** Indicates whether the term is of sort real. *)
+ val is_real : Expr.expr -> bool
+
+ (** Indicates whether the term is an integer numeral. *)
+ val is_int_numeral : Expr.expr -> bool
+
+ (** Indicates whether the term is a real numeral. *)
+ val is_rat_num : Expr.expr -> bool
+
+ (** Indicates whether the term is an algebraic number *)
+ val is_algebraic_number : Expr.expr -> bool
+
+ (** Create an expression representing t[0] + t[1] + .... *)
+ val mk_add : context -> arith_expr list -> arith_expr
+
+ (** Create an expression representing t[0] * t[1] * .... *)
+ val mk_mul : context -> arith_expr list -> arith_expr
+
+ (** Create an expression representing t[0] - t[1] - .... *)
+ val mk_sub : context -> arith_expr list -> arith_expr
+
+ (** Create an expression representing -t. *)
+ val mk_unary_minus : context -> arith_expr -> arith_expr
+
+ (** Create an expression representing t1 / t2. *)
+ val mk_div : context -> arith_expr -> arith_expr -> arith_expr
+
+ (** Create an expression representing t1 ^ t2. *)
+ val mk_power : context -> arith_expr -> arith_expr -> arith_expr
+
+ (** Create an expression representing t1 < t2 *)
+ val mk_lt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+
+ (** Create an expression representing t1 <= t2 *)
+ val mk_le : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+
+ (** Create an expression representing t1 > t2 *)
+ val mk_gt : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+
+ (** Create an expression representing t1 >= t2 *)
+ val mk_ge : context -> arith_expr -> arith_expr -> Boolean.bool_expr
+end
+
+(** Functions to manipulate bit-vector expressions *)
+and BitVector :
+sig
+ type bitvec_sort = BitVecSort of Sort.sort
+ type bitvec_expr = BitVecExpr of Expr.expr
+ type bitvec_num = BitVecNum of bitvec_expr
+
+ val sort_of_bitvec_sort : BitVector.bitvec_sort -> Sort.sort
+ val bitvec_sort_of_sort : Sort.sort -> BitVector.bitvec_sort
+ val expr_of_bitvec_expr : BitVector.bitvec_expr -> Expr.expr
+ val bitvec_expr_of_bitvec_num : BitVector.bitvec_num -> BitVector.bitvec_expr
+ val bitvec_expr_of_expr : Expr.expr -> BitVector.bitvec_expr
+ val bitvec_num_of_bitvec_expr : BitVector.bitvec_expr -> BitVector.bitvec_num
+
+ (** Create a new bit-vector sort. *)
+ val mk_sort : context -> int -> bitvec_sort
+
+ (** Indicates whether the terms is of bit-vector sort. *)
+ val is_bv : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector numeral *)
+ val is_bv_numeral : Expr.expr -> bool
+
+ (** Indicates whether the term is a one-bit bit-vector with value one *)
+ val is_bv_bit1 : Expr.expr -> bool
+
+ (** Indicates whether the term is a one-bit bit-vector with value zero *)
+ val is_bv_bit0 : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector unary minus *)
+ val is_bv_uminus : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector addition (binary) *)
+ val is_bv_add : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector subtraction (binary) *)
+ val is_bv_sub : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector multiplication (binary) *)
+ val is_bv_mul : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector signed division (binary) *)
+ val is_bv_sdiv : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector unsigned division (binary) *)
+ val is_bv_udiv : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector signed remainder (binary) *)
+ val is_bv_SRem : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector unsigned remainder (binary) *)
+ val is_bv_urem : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector signed modulus *)
+ val is_bv_smod : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector signed division by zero *)
+ val is_bv_sdiv0 : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector unsigned division by zero *)
+ val is_bv_udiv0 : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector signed remainder by zero *)
+ val is_bv_srem0 : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector unsigned remainder by zero *)
+ val is_bv_urem0 : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector signed modulus by zero *)
+ val is_bv_smod0 : Expr.expr -> bool
+
+ (** Indicates whether the term is an unsigned bit-vector less-than-or-equal *)
+ val is_bv_ule : Expr.expr -> bool
+
+ (** Indicates whether the term is a signed bit-vector less-than-or-equal *)
+ val is_bv_sle : Expr.expr -> bool
+
+ (** Indicates whether the term is an unsigned bit-vector greater-than-or-equal *)
+ val is_bv_uge : Expr.expr -> bool
+
+ (** Indicates whether the term is a signed bit-vector greater-than-or-equal *)
+ val is_bv_sge : Expr.expr -> bool
+
+ (** Indicates whether the term is an unsigned bit-vector less-than *)
+ val is_bv_ult : Expr.expr -> bool
+
+ (** Indicates whether the term is a signed bit-vector less-than *)
+ val is_bv_slt : Expr.expr -> bool
+
+ (** Indicates whether the term is an unsigned bit-vector greater-than *)
+ val is_bv_ugt : Expr.expr -> bool
+
+ (** Indicates whether the term is a signed bit-vector greater-than *)
+ val is_bv_sgt : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-wise AND *)
+ val is_bv_and : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-wise OR *)
+ val is_bv_or : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-wise NOT *)
+ val is_bv_not : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-wise XOR *)
+ val is_bv_xor : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-wise NAND *)
+ val is_bv_nand : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-wise NOR *)
+ val is_bv_nor : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-wise XNOR *)
+ val is_bv_xnor : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector concatenation (binary) *)
+ val is_bv_concat : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector sign extension *)
+ val is_bv_signextension : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector zero extension *)
+ val is_bv_zeroextension : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector extraction *)
+ val is_bv_extract : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector repetition *)
+ val is_bv_repeat : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector reduce OR *)
+ val is_bv_reduceor : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector reduce AND *)
+ val is_bv_reduceand : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector comparison *)
+ val is_bv_comp : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector shift left *)
+ val is_bv_shiftleft : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector logical shift right *)
+ val is_bv_shiftrightlogical : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector arithmetic shift left *)
+ val is_bv_shiftrightarithmetic : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector rotate left *)
+ val is_bv_rotateleft : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector rotate right *)
+ val is_bv_rotateright : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector rotate left (extended)
+ Similar to Z3_OP_ROTATE_LEFT, but it is a binary operator instead of a parametric one. *)
+ val is_bv_rotateleftextended : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector rotate right (extended)
+ Similar to Z3_OP_ROTATE_RIGHT, but it is a binary operator instead of a parametric one. *)
+ val is_bv_rotaterightextended : Expr.expr -> bool
+
+ (** Indicates whether the term is a coercion from integer to bit-vector
+ This function is not supported by the decision procedures. Only the most
+ rudimentary simplification rules are applied to this function. *)
+
+ (** Indicates whether the term is a coercion from bit-vector to integer
+ This function is not supported by the decision procedures. Only the most
+ rudimentary simplification rules are applied to this function. *)
+ val is_int_to_bv : Expr.expr -> bool
+
+ (** Indicates whether the term is a coercion from integer to bit-vector
+ This function is not supported by the decision procedures. Only the most
+ rudimentary simplification rules are applied to this function. *)
+ val is_bv_to_int : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector carry
+ Compute the carry bit in a full-adder. The meaning is given by the
+ equivalence (carry l1 l2 l3) <=> (or (and l1 l2) (and l1 l3) (and l2 l3))) *)
+ val is_bv_carry : Expr.expr -> bool
+
+ (** Indicates whether the term is a bit-vector ternary XOR
+ The meaning is given by the equivalence (xor3 l1 l2 l3) <=> (xor (xor l1 l2) l3) *)
+ val is_bv_xor3 : Expr.expr -> bool
+
+ (** The size of a bit-vector sort. *)
+ val get_size : bitvec_sort -> int
+
+ (** Retrieve the int value. *)
+ val get_int : bitvec_num -> int
+
+ (** Returns a string representation of the numeral. *)
+ val to_string : bitvec_num -> string
+
+ (** Creates a bit-vector constant. *)
+ val mk_const : context -> Symbol.symbol -> int -> bitvec_expr
+
+ (** Creates a bit-vector constant. *)
+ val mk_const_s : context -> string -> int -> bitvec_expr
+
+ (** Bitwise negation.
+ The argument must have a bit-vector sort. *)
+ val mk_not : context -> bitvec_expr -> Expr.expr
+
+ (** Take conjunction of bits in a vector,vector of length 1.
+ The argument must have a bit-vector sort. *)
+ val mk_redand : context -> bitvec_expr -> Expr.expr
+
+ (** Take disjunction of bits in a vector,vector of length 1.
+ The argument must have a bit-vector sort. *)
+ val mk_redor : context -> bitvec_expr -> Expr.expr
+
+ (** Bitwise conjunction.
+ The arguments must have a bit-vector sort. *)
+ val mk_and : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Bitwise disjunction.
+ The arguments must have a bit-vector sort. *)
+ val mk_or : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Bitwise XOR.
+ The arguments must have a bit-vector sort. *)
+ val mk_xor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Bitwise NAND.
+ The arguments must have a bit-vector sort. *)
+ val mk_nand : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Bitwise NOR.
+ The arguments must have a bit-vector sort. *)
+ val mk_nor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Bitwise XNOR.
+ The arguments must have a bit-vector sort. *)
+ val mk_xnor : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Standard two's complement unary minus.
+ The arguments must have a bit-vector sort. *)
+ val mk_neg : context -> bitvec_expr -> bitvec_expr
+
+ (** Two's complement addition.
+ The arguments must have the same bit-vector sort. *)
+ val mk_add : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Two's complement subtraction.
+ The arguments must have the same bit-vector sort. *)
+ val mk_sub : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Two's complement multiplication.
+ The arguments must have the same bit-vector sort. *)
+ val mk_mul : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Unsigned division.
+
+
+ It is defined as the floor of t1/t2 if \c t2 is
+ different from zero. If t2 is zero, then the result
+ is undefined.
+ The arguments must have the same bit-vector sort. *)
+ val mk_udiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Signed division.
+
+ It is defined in the following way:
+
+ - The \c floor of t1/t2 if \c t2 is different from zero, and t1*t2 >= 0.
+
+ - The \c ceiling of t1/t2 if \c t2 is different from zero, and t1*t2 < 0.
+
+ If t2 is zero, then the result is undefined.
+ The arguments must have the same bit-vector sort. *)
+ val mk_sdiv : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Unsigned remainder.
+
+ It is defined as t1 - (t1 /u t2) * t2, where /u represents unsigned division.
+ If t2 is zero, then the result is undefined.
+ The arguments must have the same bit-vector sort. *)
+ val mk_urem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Signed remainder.
+
+ It is defined as t1 - (t1 /s t2) * t2, where /s represents signed division.
+ The most significant bit (sign) of the result is equal to the most significant bit of \c t1.
+
+ If t2 is zero, then the result is undefined.
+ The arguments must have the same bit-vector sort. *)
+ val mk_srem : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Two's complement signed remainder (sign follows divisor).
+
+ If t2 is zero, then the result is undefined.
+ The arguments must have the same bit-vector sort. *)
+ val mk_smod : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Unsigned less-than
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_ult : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Two's complement signed less-than
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_slt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Unsigned less-than or equal to.
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_ule : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Two's complement signed less-than or equal to.
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_sle : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Unsigned greater than or equal to.
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_uge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Two's complement signed greater than or equal to.
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_sge : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Unsigned greater-than.
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_ugt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Two's complement signed greater-than.
+
+ The arguments must have the same bit-vector sort. *)
+ val mk_sgt : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Bit-vector concatenation.
+
+ The arguments must have a bit-vector sort.
+ @return
+ The result is a bit-vector of size n1+n2, where n1 (n2)
+ is the size of t1 (t2). *)
+ val mk_concat : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Bit-vector extraction.
+
+ Extract the bits between two limits from a bitvector of
+ size m to yield a new bitvector of size n, where
+ n = high - low + 1. *)
+ val mk_extract : context -> int -> int -> bitvec_expr -> bitvec_expr
+
+ (** Bit-vector sign extension.
+
+ Sign-extends the given bit-vector to the (signed) equivalent bitvector of
+ size m+i, where \c m is the size of the given bit-vector. *)
+ val mk_sign_ext : context -> int -> bitvec_expr -> bitvec_expr
+
+ (** Bit-vector zero extension.
+
+ Extend the given bit-vector with zeros to the (unsigned) equivalent
+ bitvector of size m+i, where \c m is the size of the
+ given bit-vector. *)
+ val mk_zero_ext : context -> int -> bitvec_expr -> bitvec_expr
+
+ (** Bit-vector repetition. *)
+ val mk_repeat : context -> int -> bitvec_expr -> bitvec_expr
+
+ (** Shift left.
+
+ It is equivalent to multiplication by 2^x where \c x is the value of third argument.
+
+ NB. The semantics of shift operations varies between environments. This
+ definition does not necessarily capture directly the semantics of the
+ programming language or assembly architecture you are modeling.*)
+ val mk_shl : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Logical shift right
+
+ It is equivalent to unsigned division by 2^x where \c x is the value of the third argument.
+
+ NB. The semantics of shift operations varies between environments. This
+ definition does not necessarily capture directly the semantics of the
+ programming language or assembly architecture you are modeling.
+
+ The arguments must have a bit-vector sort. *)
+ val mk_lshr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Arithmetic shift right
+
+ It is like logical shift right except that the most significant
+ bits of the result always copy the most significant bit of the
+ second argument.
+
+ NB. The semantics of shift operations varies between environments. This
+ definition does not necessarily capture directly the semantics of the
+ programming language or assembly architecture you are modeling.
+
+ The arguments must have a bit-vector sort. *)
+ val mk_ashr : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Rotate Left.
+ Rotate bits of \c t to the left \c i times. *)
+ val mk_rotate_left : context -> int -> bitvec_expr -> bitvec_expr
+
+ (** Rotate Right.
+ Rotate bits of \c t to the right \c i times.*)
+ val mk_rotate_right : context -> int -> bitvec_expr -> bitvec_expr
+
+ (** Rotate Left.
+ Rotate bits of the second argument to the left.*)
+ val mk_ext_rotate_left : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Rotate Right.
+ Rotate bits of the second argument to the right. *)
+ val mk_ext_rotate_right : context -> bitvec_expr -> bitvec_expr -> bitvec_expr
+
+ (** Create an integer from the bit-vector argument
+
+ If \c is_signed is false, then the bit-vector \c t1 is treated as unsigned.
+ So the result is non-negative and in the range [0..2^N-1], where
+ N are the number of bits in the argument.
+ If \c is_signed is true, \c t1 is treated as a signed bit-vector.
+
+ NB. This function is essentially treated as uninterpreted.
+ So you cannot expect Z3 to precisely reflect the semantics of this function
+ when solving constraints with this function.*)
+ val mk_bv2int : context -> bitvec_expr -> bool -> Arithmetic.Integer.int_expr
+
+ (** Create a predicate that checks that the bit-wise addition does not overflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_add_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+
+ (** Create a predicate that checks that the bit-wise addition does not underflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_add_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Create a predicate that checks that the bit-wise subtraction does not overflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_sub_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Create a predicate that checks that the bit-wise subtraction does not underflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_sub_no_underflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+
+ (** Create a predicate that checks that the bit-wise signed division does not overflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_sdiv_no_overflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Create a predicate that checks that the bit-wise negation does not overflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_neg_no_overflow : context -> bitvec_expr -> Boolean.bool_expr
+
+ (** Create a predicate that checks that the bit-wise multiplication does not overflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_mul_no_overflow : context -> bitvec_expr -> bitvec_expr -> bool -> Boolean.bool_expr
+
+ (** Create a predicate that checks that the bit-wise multiplication does not underflow.
+
+ The arguments must be of bit-vector sort. *)
+ val mk_mul_no_underflow : context -> bitvec_expr -> bitvec_expr -> Boolean.bool_expr
+
+ (** Create a bit-vector numeral. *)
+ val mk_numeral : context -> string -> int -> bitvec_num
+end
+
+(** Functions to manipulate proof expressions *)
+module Proof :
+sig
+ (** Indicates whether the term is a Proof for the expression 'true'. *)
+ val is_true : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for a fact asserted by the user. *)
+ val is_asserted : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for a fact (tagged as goal) asserted by the user. *)
+ val is_goal : Expr.expr -> bool
+
+ (** Indicates whether the term is proof via modus ponens
+
+ Given a proof for p and a proof for (implies p q), produces a proof for q.
+ T1: p
+ T2: (implies p q)
+ [mp T1 T2]: q
+ The second antecedents may also be a proof for (iff p q). *)
+ val is_modus_ponens : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for (R t t), where R is a reflexive relation.
+ This proof object has no antecedents.
+ The only reflexive relations that are used are
+ equivalence modulo namings, equality and equivalence.
+ That is, R is either '~', '=' or 'iff'. *)
+ val is_reflexivity : Expr.expr -> bool
+
+ (** Indicates whether the term is proof by symmetricity of a relation
+
+ Given an symmetric relation R and a proof for (R t s), produces a proof for (R s t).
+ T1: (R t s)
+ [symmetry T1]: (R s t)
+ T1 is the antecedent of this proof object. *)
+ val is_symmetry : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by transitivity of a relation
+
+ Given a transitive relation R, and proofs for (R t s) and (R s u), produces a proof
+ for (R t u).
+ T1: (R t s)
+ T2: (R s u)
+ [trans T1 T2]: (R t u) *)
+ val is_transitivity : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by condensed transitivity of a relation
+
+ Condensed transitivity proof. This proof object is only used if the parameter PROOF_MODE is 1.
+ It combines several symmetry and transitivity proofs.
+ Example:
+ T1: (R a b)
+ T2: (R c b)
+ T3: (R c d)
+ [trans* T1 T2 T3]: (R a d)
+ R must be a symmetric and transitive relation.
+
+ Assuming that this proof object is a proof for (R s t), then
+ a proof checker must check if it is possible to prove (R s t)
+ using the antecedents, symmetry and transitivity. That is,
+ if there is a path from s to t, if we view every
+ antecedent (R a b) as an edge between a and b. *)
+ val is_Transitivity_star : Expr.expr -> bool
+
+ (** Indicates whether the term is a monotonicity proof object.
+
+ T1: (R t_1 s_1)
+ ...
+ Tn: (R t_n s_n)
+ [monotonicity T1 ... Tn]: (R (f t_1 ... t_n) (f s_1 ... s_n))
+ Remark: if t_i == s_i, then the antecedent Ti is suppressed.
+ That is, reflexivity proofs are supressed to save space. *)
+ val is_monotonicity : Expr.expr -> bool
+
+ (** Indicates whether the term is a quant-intro proof
+
+ Given a proof for (~ p q), produces a proof for (~ (forall (x) p) (forall (x) q)).
+ T1: (~ p q)
+ [quant-intro T1]: (~ (forall (x) p) (forall (x) q)) *)
+ val is_quant_intro : Expr.expr -> bool
+
+ (** Indicates whether the term is a distributivity proof object.
+
+ Given that f (= or) distributes over g (= and), produces a proof for
+ (= (f a (g c d))
+ (g (f a c) (f a d)))
+ If f and g are associative, this proof also justifies the following equality:
+ (= (f (g a b) (g c d))
+ (g (f a c) (f a d) (f b c) (f b d)))
+ where each f and g can have arbitrary number of arguments.
+
+ This proof object has no antecedents.
+ Remark. This rule is used by the CNF conversion pass and
+ instantiated by f = or, and g = and. *)
+ val is_distributivity : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by elimination of AND
+
+ Given a proof for (and l_1 ... l_n), produces a proof for l_i
+ T1: (and l_1 ... l_n)
+ [and-elim T1]: l_i *)
+ val is_and_elimination : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by eliminiation of not-or
+
+ Given a proof for (not (or l_1 ... l_n)), produces a proof for (not l_i).
+ T1: (not (or l_1 ... l_n))
+ [not-or-elim T1]: (not l_i) *)
+ val is_or_elimination : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by rewriting
+
+ A proof for a local rewriting step (= t s).
+ The head function symbol of t is interpreted.
+
+ This proof object has no antecedents.
+ The conclusion of a rewrite rule is either an equality (= t s),
+ an equivalence (iff t s), or equi-satisfiability (~ t s).
+ Remark: if f is bool, then = is iff.
+
+ Examples:
+ (= (+ ( x : expr ) 0) x)
+ (= (+ ( x : expr ) 1 2) (+ 3 x))
+ (iff (or ( x : expr ) false) x) *)
+ val is_rewrite : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by rewriting
+
+ A proof for rewriting an expression t into an expression s.
+ This proof object is used if the parameter PROOF_MODE is 1.
+ This proof object can have n antecedents.
+ The antecedents are proofs for equalities used as substitution rules.
+ The object is also used in a few cases if the parameter PROOF_MODE is 2.
+ The cases are:
+ - When applying contextual simplification (CONTEXT_SIMPLIFIER=true)
+ - When converting bit-vectors to Booleans (BIT2BOOL=true)
+ - When pulling ite expression up (PULL_CHEAP_ITE_TREES=true) *)
+ val is_rewrite_star : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for pulling quantifiers out.
+
+ A proof for (iff (f (forall (x) q(x)) r) (forall (x) (f (q x) r))). This proof object has no antecedents. *)
+ val is_pull_quant : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for pulling quantifiers out.
+
+ A proof for (iff P Q) where Q is in prenex normal form.
+ This proof object is only used if the parameter PROOF_MODE is 1.
+ This proof object has no antecedents *)
+ val is_pull_quant_star : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for pushing quantifiers in.
+
+ A proof for:
+ (iff (forall (x_1 ... x_m) (and p_1[x_1 ... x_m] ... p_n[x_1 ... x_m]))
+ (and (forall (x_1 ... x_m) p_1[x_1 ... x_m])
+ ...
+ (forall (x_1 ... x_m) p_n[x_1 ... x_m])))
+ This proof object has no antecedents *)
+ val is_push_quant : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for elimination of unused variables.
+
+ A proof for (iff (forall (x_1 ... x_n y_1 ... y_m) p[x_1 ... x_n])
+ (forall (x_1 ... x_n) p[x_1 ... x_n]))
+
+ It is used to justify the elimination of unused variables.
+ This proof object has no antecedents. *)
+ val is_elim_unused_vars : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for destructive equality resolution
+
+ A proof for destructive equality resolution:
+ (iff (forall (x) (or (not (= ( x : expr ) t)) P[x])) P[t])
+ if ( x : expr ) does not occur in t.
+
+ This proof object has no antecedents.
+
+ Several variables can be eliminated simultaneously. *)
+ val is_der : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for quantifier instantiation
+
+ A proof of (or (not (forall (x) (P x))) (P a)) *)
+ val is_quant_inst : Expr.expr -> bool
+
+ (** Indicates whether the term is a hypthesis marker.
+ Mark a hypothesis in a natural deduction style proof. *)
+ val is_hypothesis : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by lemma
+
+ T1: false
+ [lemma T1]: (or (not l_1) ... (not l_n))
+
+ This proof object has one antecedent: a hypothetical proof for false.
+ It converts the proof in a proof for (or (not l_1) ... (not l_n)),
+ when T1 contains the hypotheses: l_1, ..., l_n. *)
+ val is_lemma : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by unit resolution
+
+ T1: (or l_1 ... l_n l_1' ... l_m')
+ T2: (not l_1)
+ ...
+ T(n+1): (not l_n)
+ [unit-resolution T1 ... T(n+1)]: (or l_1' ... l_m') *)
+ val is_unit_resolution : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by iff-true
+
+ T1: p
+ [iff-true T1]: (iff p true) *)
+ val is_iff_true : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by iff-false
+
+ T1: (not p)
+ [iff-false T1]: (iff p false) *)
+ val is_iff_false : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by commutativity
+
+ [comm]: (= (f a b) (f b a))
+
+ f is a commutative operator.
+
+ This proof object has no antecedents.
+ Remark: if f is bool, then = is iff. *)
+ val is_commutativity : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for Tseitin-like axioms
+
+ Proof object used to justify Tseitin's like axioms:
+
+ (or (not (and p q)) p)
+ (or (not (and p q)) q)
+ (or (not (and p q r)) p)
+ (or (not (and p q r)) q)
+ (or (not (and p q r)) r)
+ ...
+ (or (and p q) (not p) (not q))
+ (or (not (or p q)) p q)
+ (or (or p q) (not p))
+ (or (or p q) (not q))
+ (or (not (iff p q)) (not p) q)
+ (or (not (iff p q)) p (not q))
+ (or (iff p q) (not p) (not q))
+ (or (iff p q) p q)
+ (or (not (ite a b c)) (not a) b)
+ (or (not (ite a b c)) a c)
+ (or (ite a b c) (not a) (not b))
+ (or (ite a b c) a (not c))
+ (or (not (not a)) (not a))
+ (or (not a) a)
+
+ This proof object has no antecedents.
+ Note: all axioms are propositional tautologies.
+ Note also that 'and' and 'or' can take multiple arguments.
+ You can recover the propositional tautologies by
+ unfolding the Boolean connectives in the axioms a small
+ bounded number of steps (=3). *)
+ val is_def_axiom : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for introduction of a name
+
+ Introduces a name for a formula/term.
+ Suppose e is an expression with free variables x, and def-intro
+ introduces the name n(x). The possible cases are:
+
+ When e is of Boolean type:
+ [def-intro]: (and (or n (not e)) (or (not n) e))
+
+ or:
+ [def-intro]: (or (not n) e)
+ when e only occurs positively.
+
+ When e is of the form (ite cond th el):
+ [def-intro]: (and (or (not cond) (= n th)) (or cond (= n el)))
+
+ Otherwise:
+ [def-intro]: (= n e) *)
+ val is_def_intro : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for application of a definition
+
+ [apply-def T1]: F ~ n
+ F is 'equivalent' to n, given that T1 is a proof that
+ n is a name for F. *)
+ val is_apply_def : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof iff-oeq
+
+ T1: (iff p q)
+ [iff~ T1]: (~ p q) *)
+ val is_iff_oeq : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for a positive NNF step
+
+ Proof for a (positive) NNF step. Example:
+
+ T1: (not s_1) ~ r_1
+ T2: (not s_2) ~ r_2
+ T3: s_1 ~ r_1'
+ T4: s_2 ~ r_2'
+ [nnf-pos T1 T2 T3 T4]: (~ (iff s_1 s_2)
+ (and (or r_1 r_2') (or r_1' r_2)))
+
+ The negation normal form steps NNF_POS and NNF_NEG are used in the following cases:
+ (a) When creating the NNF of a positive force quantifier.
+ The quantifier is retained (unless the bound variables are eliminated).
+ Example
+ T1: q ~ q_new
+ [nnf-pos T1]: (~ (forall (x T) q) (forall (x T) q_new))
+
+ (b) When recursively creating NNF over Boolean formulas, where the top-level
+ connective is changed during NNF conversion. The relevant Boolean connectives
+ for NNF_POS are 'implies', 'iff', 'xor', 'ite'.
+ NNF_NEG furthermore handles the case where negation is pushed
+ over Boolean connectives 'and' and 'or'. *)
+ val is_nnf_pos : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for a negative NNF step
+
+ Proof for a (negative) NNF step. Examples:
+
+ T1: (not s_1) ~ r_1
+ ...
+ Tn: (not s_n) ~ r_n
+ [nnf-neg T1 ... Tn]: (not (and s_1 ... s_n)) ~ (or r_1 ... r_n)
+ and
+ T1: (not s_1) ~ r_1
+ ...
+ Tn: (not s_n) ~ r_n
+ [nnf-neg T1 ... Tn]: (not (or s_1 ... s_n)) ~ (and r_1 ... r_n)
+ and
+ T1: (not s_1) ~ r_1
+ T2: (not s_2) ~ r_2
+ T3: s_1 ~ r_1'
+ T4: s_2 ~ r_2'
+ [nnf-neg T1 T2 T3 T4]: (~ (not (iff s_1 s_2))
+ (and (or r_1 r_2) (or r_1' r_2'))) *)
+ val is_nnf_neg : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for (~ P Q) here Q is in negation normal form.
+
+ A proof for (~ P Q) where Q is in negation normal form.
+
+ This proof object is only used if the parameter PROOF_MODE is 1.
+
+ This proof object may have n antecedents. Each antecedent is a PR_DEF_INTRO. *)
+ val is_nnf_star : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for (~ P Q) where Q is in conjunctive normal form.
+
+ A proof for (~ P Q) where Q is in conjunctive normal form.
+ This proof object is only used if the parameter PROOF_MODE is 1.
+ This proof object may have n antecedents. Each antecedent is a PR_DEF_INTRO. *)
+ val is_cnf_star : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for a Skolemization step
+
+ Proof for:
+
+ [sk]: (~ (not (forall ( x : expr ) (p ( x : expr ) y))) (not (p (sk y) y)))
+ [sk]: (~ (exists ( x : expr ) (p ( x : expr ) y)) (p (sk y) y))
+
+ This proof object has no antecedents. *)
+ val is_skolemize : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof by modus ponens for equi-satisfiability.
+
+ Modus ponens style rule for equi-satisfiability.
+ T1: p
+ T2: (~ p q)
+ [mp~ T1 T2]: q *)
+ val is_modus_ponens_oeq : Expr.expr -> bool
+
+ (** Indicates whether the term is a proof for theory lemma
+
+ Generic proof for theory lemmas.
+
+ The theory lemma function comes with one or more parameters.
+ The first parameter indicates the name of the theory.
+ For the theory of arithmetic, additional parameters provide hints for
+ checking the theory lemma.
+ The hints for arithmetic are:
+ - farkas - followed by rational coefficients. Multiply the coefficients to the
+ inequalities in the lemma, add the (negated) inequalities and obtain a contradiction.
+ - triangle-eq - Indicates a lemma related to the equivalence:
+ (iff (= t1 t2) (and (<= t1 t2) (<= t2 t1)))
+ - gcd-test - Indicates an integer linear arithmetic lemma that uses a gcd test. *)
+ val is_theory_lemma : Expr.expr -> bool
+end
+
+(** Goals
+
+ A goal (aka problem). A goal is essentially a
+ of formulas, that can be solved and/or transformed using
+ tactics and solvers. *)
+module Goal :
+sig
+ type goal
+
+ (** The precision of the goal.
+
+ Goals can be transformed using over and under approximations.
+ An under approximation is applied when the objective is to find a model for a given goal.
+ An over approximation is applied when the objective is to find a proof for a given goal. *)
+ val get_precision : goal -> Z3enums.goal_prec
+
+ (** Indicates whether the goal is precise. *)
+ val is_precise : goal -> bool
+
+ (** Indicates whether the goal is an under-approximation. *)
+ val is_underapproximation : goal -> bool
+
+ (** Indicates whether the goal is an over-approximation. *)
+ val is_overapproximation : goal -> bool
+
+ (** Indicates whether the goal is garbage (i.e., the product of over- and under-approximations). *)
+ val is_garbage : goal -> bool
+
+ (** Adds the constraints to the given goal. *)
+ val assert_ : goal -> Boolean.bool_expr list -> unit
+
+ (** Indicates whether the goal contains `false'. *)
+ val is_inconsistent : goal -> bool
+
+ (** The depth of the goal.
+ This tracks how many transformations were applied to it. *)
+ val get_depth : goal -> int
+
+ (** Erases all formulas from the given goal. *)
+ val reset : goal -> unit
+
+ (** The number of formulas in the goal. *)
+ val get_size : goal -> int
+
+ (** The formulas in the goal. *)
+ val get_formulas : goal -> Boolean.bool_expr list
+
+ (** The number of formulas, subformulas and terms in the goal. *)
+ val get_num_exprs : goal -> int
+
+ (** Indicates whether the goal is empty, and it is precise or the product of an under approximation. *)
+ val is_decided_sat : goal -> bool
+
+ (** Indicates whether the goal contains `false', and it is precise or the product of an over approximation. *)
+ val is_decided_unsat : goal -> bool
+
+ (** Translates (copies) the Goal to another context.. *)
+ val translate : goal -> context -> goal
+
+ (** Simplifies the goal. Essentially invokes the `simplify' tactic on the goal. *)
+ val simplify : goal -> Params.params option -> goal
+
+ (** Creates a new Goal.
+
+ Note that the Context must have been created with proof generation support if
+ the fourth argument is set to true here. *)
+ val mk_goal : context -> bool -> bool -> bool -> goal
+
+ (** A string representation of the Goal. *)
+ val to_string : goal -> string
+end
+
+(** Models
+
+ A Model contains interpretations (assignments) of constants and functions. *)
+module Model :
+sig
+ type model
+
+ (** Function interpretations
+
+ A function interpretation is represented as a finite map and an 'else'.
+ Each entry in the finite map represents the value of a function given a set of arguments. *)
+ module FuncInterp :
+ sig
+ type func_interp
+
+ (** Function interpretations entries
+
+ An Entry object represents an element in the finite map used to a function interpretation. *)
+ module FuncEntry :
+ sig
+ type func_entry
+
+ (** Return the (symbolic) value of this entry.
+ *)
+ val get_value : func_entry -> Expr.expr
+
+ (** The number of arguments of the entry.
+ *)
+ val get_num_args : func_entry -> int
+
+ (** The arguments of the function entry.
+ *)
+ val get_args : func_entry -> Expr.expr list
+
+ (** A string representation of the function entry.
+ *)
+ val to_string : func_entry -> string
+ end
+
+ (** The number of entries in the function interpretation. *)
+ val get_num_entries : func_interp -> int
+
+ (** The entries in the function interpretation *)
+ val get_entries : func_interp -> FuncEntry.func_entry list
+
+ (** The (symbolic) `else' value of the function interpretation. *)
+ val get_else : func_interp -> Expr.expr
+
+ (** The arity of the function interpretation *)
+ val get_arity : func_interp -> int
+
+ (** A string representation of the function interpretation. *)
+ val to_string : func_interp -> string
+ end
+
+ (** Retrieves the interpretation (the assignment) of a func_decl in the model.
+ An expression if the function has an interpretation in the model, null otherwise. *)
+ val get_const_interp : model -> FuncDecl.func_decl -> Expr.expr option
+
+ (** Retrieves the interpretation (the assignment) of an expression in the model.
+ An expression if the constant has an interpretation in the model, null otherwise. *)
+ val get_const_interp_e : model -> Expr.expr -> Expr.expr option
+
+ (** Retrieves the interpretation (the assignment) of a non-constant func_decl in the model.
+ A FunctionInterpretation if the function has an interpretation in the model, null otherwise. *)
+ val get_func_interp : model -> FuncDecl.func_decl -> FuncInterp.func_interp option
+
+ (** The number of constant interpretations in the model. *)
+ val get_num_consts : model -> int
+
+ (** The function declarations of the constants in the model. *)
+ val get_const_decls : model -> FuncDecl.func_decl list
+
+ (** The number of function interpretations in the model. *)
+ val get_num_funcs : model -> int
+
+ (** The function declarations of the function interpretations in the model. *)
+ val get_func_decls : model -> FuncDecl.func_decl list
+
+ (** All symbols that have an interpretation in the model. *)
+ val get_decls : model -> FuncDecl.func_decl list
+
+ (** A ModelEvaluationFailedException is thrown when an expression cannot be evaluated by the model. *)
+ exception ModelEvaluationFailedException of string
+
+ (** Evaluates an expression in the current model.
+
+ This function may fail if the argument contains quantifiers,
+ is partial (MODEL_PARTIAL enabled), or if it is not well-sorted.
+ In this case a ModelEvaluationFailedException is thrown.
+ *)
+ val eval : model -> Expr.expr -> bool -> Expr.expr
+
+ (** Alias for eval. *)
+ val evaluate : model -> Expr.expr -> bool -> Expr.expr
+
+ (** The number of uninterpreted sorts that the model has an interpretation for. *)
+ val get_num_sorts : model -> int
+
+ (** The uninterpreted sorts that the model has an interpretation for.
+
+ Z3 also provides an intepretation for uninterpreted sorts used in a formula.
+ The interpretation for a sort is a finite set of distinct values. We say this finite set is
+ the "universe" of the sort.
+ {!get_num_sorts}
+ {!sort_universe} *)
+ val get_sorts : model -> Sort.sort list
+
+ (** The finite set of distinct values that represent the interpretation of a sort.
+ {!get_sorts}
+ @returns A list of expressions, where each is an element of the universe of the sort *)
+ val sort_universe : model -> Sort.sort -> AST.ast list
+
+ (** Conversion of models to strings.
+ A string representation of the model. *)
+ val to_string : model -> string
+end
+
+(** Probes
+
+ Probes are used to inspect a goal (aka problem) and collect information that may be used to decide
+ which solver and/or preprocessing step will be used.
+ The complete list of probes may be obtained using the procedures Context.NumProbes
+ and Context.ProbeNames.
+ It may also be obtained using the command (help-tactics) in the SMT 2.0 front-end.
+*)
+module Probe :
+sig
+ type probe
+
+ (** Execute the probe over the goal.
+ A probe always produce a double value.
+ "Boolean" probes return 0.0 for false, and a value different from 0.0 for true. *)
+ val apply : probe -> Goal.goal -> float
+
+ (** The number of supported Probes. *)
+ val get_num_probes : context -> int
+
+ (** The names of all supported Probes. *)
+ val get_probe_names : context -> string list
+
+ (** Returns a string containing a description of the probe with the given name. *)
+ val get_probe_description : context -> string -> string
+
+ (** Creates a new Probe. *)
+ val mk_probe : context -> string -> probe
+
+ (** Create a probe that always evaluates to a float value. *)
+ val const : context -> float -> probe
+
+ (** Create a probe that evaluates to "true" when the value returned by the first argument
+ is less than the value returned by second argument *)
+ val lt : context -> probe -> probe -> probe
+
+ (** Create a probe that evaluates to "true" when the value returned by the first argument
+ is greater than the value returned by second argument *)
+ val gt : context -> probe -> probe -> probe
+
+ (** Create a probe that evaluates to "true" when the value returned by the first argument
+ is less than or equal the value returned by second argument *)
+ val le : context -> probe -> probe -> probe
+
+ (** Create a probe that evaluates to "true" when the value returned by the first argument
+ is greater than or equal the value returned by second argument *)
+ val ge : context -> probe -> probe -> probe
+
+
+ (** Create a probe that evaluates to "true" when the value returned by the first argument
+ is equal the value returned by second argument *)
+ val eq : context -> probe -> probe -> probe
+
+ (** Create a probe that evaluates to "true" when both of two probes evaluate to "true". *)
+ val and_ : context -> probe -> probe -> probe
+
+ (** Create a probe that evaluates to "true" when either of two probes evaluates to "true". *)
+ val or_ : context -> probe -> probe -> probe
+
+ (** Create a probe that evaluates to "true" when another probe does not evaluate to "true". *)
+ val not_ : context -> probe -> probe
+end
+
+(** Tactics
+
+ Tactics are the basic building block for creating custom solvers for specific problem domains.
+ The complete list of tactics may be obtained using Context.get_num_tactics
+ and Context.get_tactic_names.
+ It may also be obtained using the command (help-tactics) in the SMT 2.0 front-end.
+*)
+module Tactic :
+sig
+ type tactic
+
+ (** Tactic application results
+
+ ApplyResult objects represent the result of an application of a
+ tactic to a goal. It contains the subgoals that were produced. *)
+ module ApplyResult :
+ sig
+ type apply_result
+
+ (** The number of Subgoals. *)
+ val get_num_subgoals : apply_result -> int
+
+ (** Retrieves the subgoals from the apply_result. *)
+ val get_subgoals : apply_result -> Goal.goal list
+
+ (** Retrieves a subgoal from the apply_result. *)
+ val get_subgoal : apply_result -> int -> Goal.goal
+
+ (** Convert a model for a subgoal into a model for the original
+ goal g, that the ApplyResult was obtained from.
+ #return A model for g *)
+ val convert_model : apply_result -> int -> Model.model -> Model.model
+
+ (** A string representation of the ApplyResult. *)
+ val to_string : apply_result -> string
+ end
+
+ (** A string containing a description of parameters accepted by the tactic. *)
+ val get_help : tactic -> string
+
+ (** Retrieves parameter descriptions for Tactics. *)
+ val get_param_descrs : tactic -> Params.ParamDescrs.param_descrs
+
+ (** Apply the tactic to the goal. *)
+ val apply : tactic -> Goal.goal -> Params.params option -> ApplyResult.apply_result
+
+ (** The number of supported tactics. *)
+ val get_num_tactics : context -> int
+
+ (** The names of all supported tactics. *)
+ val get_tactic_names : context -> string list
+
+ (** Returns a string containing a description of the tactic with the given name. *)
+ val get_tactic_description : context -> string -> string
+
+ (** Creates a new Tactic. *)
+ val mk_tactic : context -> string -> tactic
+
+ (** Create a tactic that applies one tactic to a Goal and
+ then another one to every subgoal produced by the first one. *)
+ val and_then : context -> tactic -> tactic -> tactic list -> tactic
+
+ (** Create a tactic that first applies one tactic to a Goal and
+ if it fails then returns the result of another tactic applied to the Goal. *)
+ val or_else : context -> tactic -> tactic -> tactic
+
+ (** Create a tactic that applies one tactic to a goal for some time (in milliseconds).
+
+ If the tactic does not terminate within the timeout, then it fails. *)
+ val try_for : context -> tactic -> int -> tactic
+
+ (** Create a tactic that applies one tactic to a given goal if the probe
+ evaluates to true.
+
+ If the probe evaluates to false, then the new tactic behaves like the skip tactic. *)
+ val when_ : context -> Probe.probe -> tactic -> tactic
+
+ (** Create a tactic that applies a tactic to a given goal if the probe
+ evaluates to true and another tactic otherwise. *)
+ val cond : context -> Probe.probe -> tactic -> tactic -> tactic
+
+ (** Create a tactic that keeps applying one tactic until the goal is not
+ modified anymore or the maximum number of iterations is reached. *)
+ val repeat : context -> tactic -> int -> tactic
+
+ (** Create a tactic that just returns the given goal. *)
+ val skip : context -> tactic
+
+ (** Create a tactic always fails. *)
+ val fail : context -> tactic
+
+ (** Create a tactic that fails if the probe evaluates to false. *)
+ val fail_if : context -> Probe.probe -> tactic
+
+ (** Create a tactic that fails if the goal is not triviall satisfiable (i.e., empty)
+ or trivially unsatisfiable (i.e., contains `false'). *)
+ val fail_if_not_decided : context -> tactic
+
+ (** Create a tactic that applies a tactic using the given set of parameters. *)
+ val using_params : context -> tactic -> Params.params -> tactic
+
+ (** Create a tactic that applies a tactic using the given set of parameters.
+ Alias for UsingParams*)
+ val with_ : context -> tactic -> Params.params -> tactic
+
+ (** Create a tactic that applies the given tactics in parallel. *)
+ val par_or : context -> tactic list -> tactic
+
+ (** Create a tactic that applies a tactic to a given goal and then another tactic
+ to every subgoal produced by the first one. The subgoals are processed in parallel. *)
+ val par_and_then : context -> tactic -> tactic -> tactic
+
+ (** Interrupt the execution of a Z3 procedure.
+ This procedure can be used to interrupt: solvers, simplifiers and tactics. *)
+ val interrupt : context -> unit
+end
+
+(** Solvers *)
+module Solver :
+sig
+ type solver
+ type status = UNSATISFIABLE | UNKNOWN | SATISFIABLE
+
+ val string_of_status : status -> string
+
+ (** Objects that track statistical information about solvers. *)
+ module Statistics :
+ sig
+ type statistics
+
+ (** Statistical data is organized into pairs of \[Key, Entry\], where every
+ Entry is either a floating point or integer value.
+ *)
+ module Entry :
+ sig
+ type statistics_entry
+
+ (** The key of the entry. *)
+ val get_key : statistics_entry -> string
+
+ (** The int-value of the entry. *)
+ val get_int : statistics_entry -> int
+
+ (** The float-value of the entry. *)
+ val get_float : statistics_entry -> float
+
+ (** True if the entry is uint-valued. *)
+ val is_int : statistics_entry -> bool
+
+ (** True if the entry is double-valued. *)
+ val is_float : statistics_entry -> bool
+
+ (** The string representation of the the entry's value. *)
+ val to_string_value : statistics_entry -> string
+
+ (** The string representation of the entry (key and value) *)
+ val to_string : statistics_entry -> string
+ end
+
+ (** A string representation of the statistical data. *)
+ val to_string : statistics -> string
+
+ (** The number of statistical data. *)
+ val get_size : statistics -> int
+
+ (** The data entries. *)
+ val get_entries : statistics -> Entry.statistics_entry list
+
+ (** The statistical counters. *)
+ val get_keys : statistics -> string list
+
+ (** The value of a particular statistical counter. *)
+ val get : statistics -> string -> Entry.statistics_entry option
+ end
+
+ (** A string that describes all available solver parameters. *)
+ val get_help : solver -> string
+
+ (** Sets the solver parameters. *)
+ val set_parameters : solver -> Params.params -> unit
+
+ (** Retrieves parameter descriptions for solver. *)
+ val get_param_descrs : solver -> Params.ParamDescrs.param_descrs
+
+ (** The current number of backtracking points (scopes).
+ {!pop}
+ {!push} *)
+ val get_num_scopes : solver -> int
+
+ (** Creates a backtracking point.
+ {!pop} *)
+ val push : solver -> unit
+
+ (** Backtracks a number of backtracking points.
+ Note that an exception is thrown if the integer is not smaller than {!get_num_scopes}
+ {!push} *)
+ val pop : solver -> int -> unit
+
+ (** Resets the Solver.
+ This removes all assertions from the solver. *)
+ val reset : solver -> unit
+
+ (** Assert a constraint (or multiple) into the solver. *)
+ val assert_ : solver -> Boolean.bool_expr list -> unit
+
+ (** * Assert multiple constraints (cs) into the solver, and track them (in the
+ * unsat) core
+ * using the Boolean constants in ps.
+ *
+ * This API is an alternative to {!check} with assumptions for
+ * extracting unsat cores.
+ * Both APIs can be used in the same solver. The unsat core will contain a
+ * combination
+ * of the Boolean variables provided using {!assert_and_track}
+ * and the Boolean literals
+ * provided using {!check} with assumptions. *)
+ val assert_and_track_a : solver -> Boolean.bool_expr list -> Boolean.bool_expr list -> unit
+
+ (** * Assert a constraint (c) into the solver, and track it (in the unsat) core
+ * using the Boolean constant p.
+ *
+ * This API is an alternative to {!check} with assumptions for
+ * extracting unsat cores.
+ * Both APIs can be used in the same solver. The unsat core will contain a
+ * combination
+ * of the Boolean variables provided using {!assert_and_track}
+ * and the Boolean literals
+ * provided using {!check} with assumptions. *)
+ val assert_and_track : solver -> Boolean.bool_expr -> Boolean.bool_expr -> unit
+
+ (** The number of assertions in the solver. *)
+ val get_num_assertions : solver -> int
+
+ (** The set of asserted formulas. *)
+ val get_assertions : solver -> Boolean.bool_expr list
+
+ (** Checks whether the assertions in the solver are consistent or not.
+
+ {!Model}
+ {!get_unsat_core}
+ {!Proof} *)
+ val check : solver -> Boolean.bool_expr list -> status
+
+ (** The model of the last Check.
+
+ The result is None if Check was not invoked before,
+ if its results was not SATISFIABLE, or if model production is not enabled. *)
+ val get_model : solver -> Model.model option
+
+ (** The proof of the last Check.
+
+ The result is null if Check was not invoked before,
+ if its results was not UNSATISFIABLE, or if proof production is disabled. *)
+ val get_proof : solver -> Expr.expr option
+
+ (** The unsat core of the last Check.
+
+ The unsat core is a subset of Assertions
+ The result is empty if Check was not invoked before,
+ if its results was not UNSATISFIABLE, or if core production is disabled. *)
+ val get_unsat_core : solver -> AST.ast list
+
+ (** A brief justification of why the last call to Check returned UNKNOWN. *)
+ val get_reason_unknown : solver -> string
+
+ (** Solver statistics. *)
+ val get_statistics : solver -> Statistics.statistics
+
+ (** Creates a new (incremental) solver.
+
+ This solver also uses a set of builtin tactics for handling the first
+ check-sat command, and check-sat commands that take more than a given
+ number of milliseconds to be solved. *)
+ val mk_solver : context -> Symbol.symbol option -> solver
+
+ (** Creates a new (incremental) solver.
+ {!mk_solver} *)
+ val mk_solver_s : context -> string -> solver
+
+ (** Creates a new (incremental) solver. *)
+ val mk_simple_solver : context -> solver
+
+ (** Creates a solver that is implemented using the given tactic.
+
+ The solver supports the commands Push and Pop, but it
+ will always solve each check from scratch. *)
+ val mk_solver_t : context -> Tactic.tactic -> solver
+
+ (** A string representation of the solver. *)
+ val to_string : solver -> string
+end
+
+(** Fixedpoint solving *)
+module Fixedpoint :
+sig
+ type fixedpoint
+
+ (** A string that describes all available fixedpoint solver parameters. *)
+ val get_help : fixedpoint -> string
+
+ (** Sets the fixedpoint solver parameters. *)
+ val set_params : fixedpoint -> Params.params -> unit
+
+ (** Retrieves parameter descriptions for Fixedpoint solver. *)
+ val get_param_descrs : fixedpoint -> Params.ParamDescrs.param_descrs
+
+ (** Assert a constraints into the fixedpoint solver. *)
+ val assert_ : fixedpoint -> Boolean.bool_expr list -> unit
+
+ (** Register predicate as recursive relation. *)
+ val register_relation : fixedpoint -> FuncDecl.func_decl -> unit
+
+ (** Add rule into the fixedpoint solver. *)
+ val add_rule : fixedpoint -> Boolean.bool_expr -> Symbol.symbol option -> unit
+
+ (** Add table fact to the fixedpoint solver. *)
+ val add_fact : fixedpoint -> FuncDecl.func_decl -> int list -> unit
+
+ (** Query the fixedpoint solver.
+ A query is a conjunction of constraints. The constraints may include the recursively defined relations.
+ The query is satisfiable if there is an instance of the query variables and a derivation for it.
+ The query is unsatisfiable if there are no derivations satisfying the query variables. *)
+ val query : fixedpoint -> Boolean.bool_expr -> Solver.status
+
+ (** Query the fixedpoint solver.
+ A query is an array of relations.
+ The query is satisfiable if there is an instance of some relation that is non-empty.
+ The query is unsatisfiable if there are no derivations satisfying any of the relations. *)
+ val query_r : fixedpoint -> FuncDecl.func_decl list -> Solver.status
+
+ (** Creates a backtracking point.
+ {!pop} *)
+ val push : fixedpoint -> unit
+
+ (** Backtrack one backtracking point.
+
+ Note that an exception is thrown if Pop is called without a corresponding Push
+ {!push} *)
+ val pop : fixedpoint -> unit
+
+ (** Update named rule into in the fixedpoint solver. *)
+ val update_rule : fixedpoint -> Boolean.bool_expr -> Symbol.symbol -> unit
+
+ (** Retrieve satisfying instance or instances of solver,
+ or definitions for the recursive predicates that show unsatisfiability. *)
+ val get_answer : fixedpoint -> Expr.expr option
+
+ (** Retrieve explanation why fixedpoint engine returned status Unknown. *)
+ val get_reason_unknown : fixedpoint -> string
+
+ (** Retrieve the number of levels explored for a given predicate. *)
+ val get_num_levels : fixedpoint -> FuncDecl.func_decl -> int
+
+ (** Retrieve the cover of a predicate. *)
+ val get_cover_delta : fixedpoint -> int -> FuncDecl.func_decl -> Expr.expr option
+
+ (** Add property about the predicate.
+ The property is added at level. *)
+ val add_cover : fixedpoint -> int -> FuncDecl.func_decl -> Expr.expr -> unit
+
+ (** Retrieve internal string representation of fixedpoint object. *)
+ val to_string : fixedpoint -> string
+
+ (** Instrument the Datalog engine on which table representation to use for recursive predicate. *)
+ val set_predicate_representation : fixedpoint -> FuncDecl.func_decl -> Symbol.symbol list -> unit
+
+ (** Convert benchmark given as set of axioms, rules and queries to a string. *)
+ val to_string_q : fixedpoint -> Boolean.bool_expr list -> string
+
+ (** Retrieve set of rules added to fixedpoint context. *)
+ val get_rules : fixedpoint -> Boolean.bool_expr list
+
+ (** Retrieve set of assertions added to fixedpoint context. *)
+ val get_assertions : fixedpoint -> Boolean.bool_expr list
+
+ (** Create a Fixedpoint context. *)
+ val mk_fixedpoint : context -> fixedpoint
+end
+
+(** Global and context options
+
+ Note: This module contains functions that set parameters/options for context
+ objects as well as functions that set options that are used globally, across
+ contexts.*)
+module Options :
+sig
+ (** Update a mutable configuration parameter.
+
+ The list of all configuration parameters can be obtained using the Z3 executable:
+ z3.exe -ini?
+ Only a few configuration parameters are mutable once the context is created.
+ An exception is thrown when trying to modify an immutable parameter.
+ {!get_param_value} *)
+ val update_param_value : context -> string -> string -> unit
+
+ (** Get a configuration parameter.
+
+ Returns None if the parameter value does not exist.
+ {!update_param_value} *)
+ val get_param_value : context -> string -> string option
+
+ (** Selects the format used for pretty-printing expressions.
+
+ The default mode for pretty printing expressions is to produce
+ SMT-LIB style output where common subexpressions are printed
+ at each occurrence. The mode is called PRINT_SMTLIB_FULL.
+ To print shared common subexpressions only once,
+ use the PRINT_LOW_LEVEL mode.
+ To print in way that conforms to SMT-LIB standards and uses let
+ expressions to share common sub-expressions use PRINT_SMTLIB_COMPLIANT.
+ {!AST.to_string}
+ {!Quantifier.Pattern.to_string}
+ {!FuncDecl.to_string}
+ {!Sort.to_string} *)
+ val set_print_mode : context -> Z3enums.ast_print_mode -> unit
+
+ (** Enable/disable printing of warning messages to the console.
+
+ Note that this function is static and effects the behaviour of
+ all contexts globally. *)
+ val toggle_warning_messages : bool -> unit
+end
+
+(** Functions for handling SMT and SMT2 expressions and files *)
+module SMT :
+sig
+ (** Convert a benchmark into an SMT-LIB formatted string.
+
+ @return A string representation of the benchmark. *)
+ val benchmark_to_smtstring : context -> string -> string -> string -> string -> Boolean.bool_expr list -> Boolean.bool_expr -> string
+
+ (** Parse the given string using the SMT-LIB parser.
+
+ The symbol table of the parser can be initialized using the given sorts and declarations.
+ The symbols in the arrays in the third and fifth argument
+ don't need to match the names of the sorts and declarations in the arrays in the fourth
+ and sixth argument. This is a useful feature since we can use arbitrary names to
+ reference sorts and declarations. *)
+ val parse_smtlib_string : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> unit
+
+ (** Parse the given file using the SMT-LIB parser.
+ {!parse_smtlib_string} *)
+ val parse_smtlib_file : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> unit
+
+ (** The number of SMTLIB formulas parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_num_smtlib_formulas : context -> int
+
+ (** The formulas parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_smtlib_formulas : context -> Boolean.bool_expr list
+
+ (** The number of SMTLIB assumptions parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_num_smtlib_assumptions : context -> int
+
+ (** The assumptions parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_smtlib_assumptions : context -> Boolean.bool_expr list
+
+ (** The number of SMTLIB declarations parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_num_smtlib_decls : context -> int
+
+ (** The declarations parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_smtlib_decls : context -> FuncDecl.func_decl list
+
+ (** The number of SMTLIB sorts parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_num_smtlib_sorts : context -> int
+
+ (** The sort declarations parsed by the last call to ParseSMTLIBString or ParseSMTLIBFile. *)
+ val get_smtlib_sorts : context -> Sort.sort list
+
+ (** Parse the given string using the SMT-LIB2 parser.
+
+ {!parse_smtlib_string}
+ @return A conjunction of assertions in the scope (up to push/pop) at the end of the string. *)
+ val parse_smtlib2_string : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> Boolean.bool_expr
+
+ (** Parse the given file using the SMT-LIB2 parser.
+ {!parse_smtlib2_string} *)
+ val parse_smtlib2_file : context -> string -> Symbol.symbol list -> Sort.sort list -> Symbol.symbol list -> FuncDecl.func_decl list -> Boolean.bool_expr
+end
+
+(** Set a global (or module) parameter, which is shared by all Z3 contexts.
+
+ When a Z3 module is initialized it will use the value of these parameters
+ when Z3_params objects are not provided.
+ The name of parameter can be composed of characters [a-z][A-Z], digits [0-9], '-' and '_'.
+ The character '.' is a delimiter (more later).
+ The parameter names are case-insensitive. The character '-' should be viewed as an "alias" for '_'.
+ Thus, the following parameter names are considered equivalent: "pp.decimal-precision" and "PP.DECIMAL_PRECISION".
+ This function can be used to set parameters for a specific Z3 module.
+ This can be done by using ..
+ For example:
+ (set_global_param "pp.decimal" "true")
+ will set the parameter "decimal" in the module "pp" to true.
+*)
+val set_global_param : string -> string -> unit
+
+(** Get a global (or module) parameter.
+
+ Returns None if the parameter does not exist.
+ The caller must invoke #Z3_global_param_del_value to delete the value returned at \c param_value.
+ This function cannot be invoked simultaneously from different threads without synchronization.
+ The result string stored in param_value is stored in a shared location.
+*)
+val get_global_param : string -> string option
+
+(** Restore the value of all global (and module) parameters.
+
+ This command will not affect already created objects (such as tactics and solvers)
+ {!set_global_param}
+*)
+val global_param_reset_all : unit