3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-04-06 17:44:08 +00:00

Add accessors for RCF numeral internals (#7013)

This commit is contained in:
Christoph M. Wintersteiger 2023-11-23 16:54:23 +00:00 committed by GitHub
parent 9382b96a32
commit 16753e43f1
No known key found for this signature in database
GPG key ID: 4AEE18F83AFDEB23
7 changed files with 845 additions and 117 deletions

View file

@ -1,4 +1,4 @@
(* (*
Copyright (C) 2012 Microsoft Corporation Copyright (C) 2012 Microsoft Corporation
Author: CM Wintersteiger (cwinter) 2012-12-17 Author: CM Wintersteiger (cwinter) 2012-12-17
*) *)
@ -19,7 +19,6 @@ open Z3.Arithmetic.Integer
open Z3.Arithmetic.Real open Z3.Arithmetic.Real
open Z3.BitVector open Z3.BitVector
exception TestFailedException of string exception TestFailedException of string
(** (**
@ -31,14 +30,14 @@ let model_converter_test ( ctx : context ) =
let yr = (Expr.mk_const ctx (Symbol.mk_string ctx "y") (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 let g4 = (mk_goal ctx true false false ) in
(Goal.add g4 [ (mk_gt ctx xr (Real.mk_numeral_nd ctx 10 1)) ]) ; (Goal.add g4 [ (mk_gt ctx xr (Real.mk_numeral_nd ctx 10 1)) ]) ;
(Goal.add g4 [ (mk_eq ctx (Goal.add g4 [ (mk_eq ctx
yr yr
(Arithmetic.mk_add ctx [ xr; (Real.mk_numeral_nd ctx 1 1) ])) ]) ; (Arithmetic.mk_add ctx [ xr; (Real.mk_numeral_nd ctx 1 1) ])) ]) ;
(Goal.add g4 [ (mk_gt ctx yr (Real.mk_numeral_nd ctx 1 1)) ]) ; (Goal.add g4 [ (mk_gt ctx yr (Real.mk_numeral_nd ctx 1 1)) ]) ;
( (
let ar = (Tactic.apply (mk_tactic ctx "simplify") g4 None) in let ar = (Tactic.apply (mk_tactic ctx "simplify") g4 None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
((is_decided_sat (get_subgoal ar 0)) || ((is_decided_sat (get_subgoal ar 0)) ||
(is_decided_unsat (get_subgoal ar 0)))) then (is_decided_unsat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
@ -46,8 +45,8 @@ let model_converter_test ( ctx : context ) =
); );
( (
let ar = (Tactic.apply (and_then ctx (mk_tactic ctx ("simplify")) (mk_tactic ctx "solve-eqs") []) g4 None) in let ar = (Tactic.apply (and_then ctx (mk_tactic ctx ("simplify")) (mk_tactic ctx "solve-eqs") []) g4 None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
((is_decided_sat (get_subgoal ar 0)) || ((is_decided_sat (get_subgoal ar 0)) ||
(is_decided_unsat (get_subgoal ar 0)))) then (is_decided_unsat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
@ -57,15 +56,15 @@ let model_converter_test ( ctx : context ) =
let f e = (Solver.add solver [ e ]) in let f e = (Solver.add solver [ e ]) in
ignore (List.map f (get_formulas (get_subgoal ar 0))) ; ignore (List.map f (get_formulas (get_subgoal ar 0))) ;
let q = (check solver []) in let q = (check solver []) in
if q != SATISFIABLE then if q != SATISFIABLE then
raise (TestFailedException "") raise (TestFailedException "")
else else
let m = (get_model solver) in let m = (get_model solver) in
match m with match m with
| None -> raise (TestFailedException "") | None -> raise (TestFailedException "")
| Some (m) -> | Some (m) ->
Printf.printf "Solver says: %s\n" (string_of_status q) ; Printf.printf "Solver says: %s\n" (string_of_status q) ;
Printf.printf "Model: \n%s\n" (Model.to_string m) Printf.printf "Model: \n%s\n" (Model.to_string m)
) )
(** (**
@ -79,7 +78,7 @@ let basic_tests ( ctx : context ) =
let bs = (Boolean.mk_sort ctx) in let bs = (Boolean.mk_sort ctx) in
let domain = [ bs; bs ] in let domain = [ bs; bs ] in
let f = (FuncDecl.mk_func_decl ctx fname domain bs) in let f = (FuncDecl.mk_func_decl ctx fname domain bs) in
let fapp = (mk_app ctx f let fapp = (mk_app ctx f
[ (Expr.mk_const ctx x bs); (Expr.mk_const ctx y bs) ]) in [ (Expr.mk_const ctx x bs); (Expr.mk_const ctx y bs) ]) in
let fargs2 = [ (mk_fresh_const ctx "cp" bs) ] in let fargs2 = [ (mk_fresh_const ctx "cp" bs) ] in
let domain2 = [ bs ] in let domain2 = [ bs ] in
@ -100,8 +99,8 @@ let basic_tests ( ctx : context ) =
); );
( (
let ar = (Tactic.apply (mk_tactic ctx "simplify") g None) in let ar = (Tactic.apply (mk_tactic ctx "simplify") g None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
((is_decided_sat (get_subgoal ar 0)) || ((is_decided_sat (get_subgoal ar 0)) ||
(is_decided_unsat (get_subgoal ar 0)))) then (is_decided_unsat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
@ -109,28 +108,28 @@ let basic_tests ( ctx : context ) =
); );
( (
let ar = (Tactic.apply (mk_tactic ctx "smt") g None) in let ar = (Tactic.apply (mk_tactic ctx "smt") g None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
(not (is_decided_sat (get_subgoal ar 0)))) then (not (is_decided_sat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
); );
(Goal.add g [ (mk_eq ctx (Goal.add g [ (mk_eq ctx
(mk_numeral_int ctx 1 (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))) ] ) (mk_numeral_int ctx 2 (BitVector.mk_sort ctx 32))) ] )
; ;
( (
let ar = (Tactic.apply (mk_tactic ctx "smt") g None) in let ar = (Tactic.apply (mk_tactic ctx "smt") g None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
(not (is_decided_unsat (get_subgoal ar 0)))) then (not (is_decided_unsat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
); );
( (
let g2 = (mk_goal ctx true true false) in let g2 = (mk_goal ctx true true false) in
let ar = (Tactic.apply (mk_tactic ctx "smt") g2 None) in let ar = (Tactic.apply (mk_tactic ctx "smt") g2 None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
(not (is_decided_sat (get_subgoal ar 0)))) then (not (is_decided_sat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
@ -140,10 +139,10 @@ let basic_tests ( ctx : context ) =
let g2 = (mk_goal ctx true true false) in let g2 = (mk_goal ctx true true false) in
(Goal.add g2 [ (Boolean.mk_false ctx) ]) ; (Goal.add g2 [ (Boolean.mk_false ctx) ]) ;
let ar = (Tactic.apply (mk_tactic ctx "smt") g2 None) in let ar = (Tactic.apply (mk_tactic ctx "smt") g2 None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
(not (is_decided_unsat (get_subgoal ar 0)))) then (not (is_decided_unsat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
); );
( (
@ -155,10 +154,10 @@ let basic_tests ( ctx : context ) =
let constr = (mk_eq ctx xc yc) in let constr = (mk_eq ctx xc yc) in
(Goal.add g3 [ constr ] ) ; (Goal.add g3 [ constr ] ) ;
let ar = (Tactic.apply (mk_tactic ctx "smt") g3 None) in let ar = (Tactic.apply (mk_tactic ctx "smt") g3 None) in
if ((get_num_subgoals ar) == 1 && if ((get_num_subgoals ar) == 1 &&
(not (is_decided_unsat (get_subgoal ar 0)))) then (not (is_decided_unsat (get_subgoal ar 0)))) then
raise (TestFailedException "") raise (TestFailedException "")
else else
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
) ; ) ;
model_converter_test ctx ; model_converter_test ctx ;
@ -169,12 +168,12 @@ let basic_tests ( ctx : context ) =
Printf.printf "Numerator: %s Denominator: %s\n" (Real.numeral_to_string inum) (Real.numeral_to_string iden) ; Printf.printf "Numerator: %s Denominator: %s\n" (Real.numeral_to_string inum) (Real.numeral_to_string iden) ;
if ((Real.numeral_to_string inum) <> "42" || (Real.numeral_to_string iden) <> "43") then if ((Real.numeral_to_string inum) <> "42" || (Real.numeral_to_string iden) <> "43") then
raise (TestFailedException "") raise (TestFailedException "")
else else
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
; ;
if ((to_decimal_string rn 3) <> "0.976?") then if ((to_decimal_string rn 3) <> "0.976?") then
raise (TestFailedException "") raise (TestFailedException "")
else else
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
; ;
if (to_decimal_string (Real.mk_numeral_s ctx "-1231231232/234234333") 5 <> "-5.25640?") then if (to_decimal_string (Real.mk_numeral_s ctx "-1231231232/234234333") 5 <> "-5.25640?") then
@ -193,7 +192,7 @@ let basic_tests ( ctx : context ) =
raise (TestFailedException "check") raise (TestFailedException "check")
) )
with Z3.Error(_) -> ( with Z3.Error(_) -> (
Printf.printf "Exception caught, OK.\n" Printf.printf "Exception caught, OK.\n"
) )
(** (**
@ -212,22 +211,22 @@ let quantifier_example1 ( ctx : context ) =
let xs = [ (Integer.mk_const ctx (List.nth names 0)); let xs = [ (Integer.mk_const ctx (List.nth names 0));
(Integer.mk_const ctx (List.nth names 1)); (Integer.mk_const ctx (List.nth names 1));
(Integer.mk_const ctx (List.nth names 2)) ] in (Integer.mk_const ctx (List.nth names 2)) ] in
let body_vars = (Boolean.mk_and ctx let body_vars = (Boolean.mk_and ctx
[ (mk_eq ctx [ (mk_eq ctx
(Arithmetic.mk_add ctx [ (List.nth vars 0) ; (Integer.mk_numeral_i ctx 1)]) (Arithmetic.mk_add ctx [ (List.nth vars 0) ; (Integer.mk_numeral_i ctx 1)])
(Integer.mk_numeral_i ctx 2)) ; (Integer.mk_numeral_i ctx 2)) ;
(mk_eq ctx (mk_eq ctx
(Arithmetic.mk_add ctx [ (List.nth vars 1); (Integer.mk_numeral_i ctx 2)]) (Arithmetic.mk_add ctx [ (List.nth vars 1); (Integer.mk_numeral_i ctx 2)])
(Arithmetic.mk_add ctx [ (List.nth vars 2); (Integer.mk_numeral_i ctx 3)])) ]) in (Arithmetic.mk_add ctx [ (List.nth vars 2); (Integer.mk_numeral_i ctx 3)])) ]) in
let body_const = (Boolean.mk_and ctx let body_const = (Boolean.mk_and ctx
[ (mk_eq ctx [ (mk_eq ctx
(Arithmetic.mk_add ctx [ (List.nth xs 0); (Integer.mk_numeral_i ctx 1)]) (Arithmetic.mk_add ctx [ (List.nth xs 0); (Integer.mk_numeral_i ctx 1)])
(Integer.mk_numeral_i ctx 2)) ; (Integer.mk_numeral_i ctx 2)) ;
(mk_eq ctx (mk_eq ctx
(Arithmetic.mk_add ctx [ (List.nth xs 1); (Integer.mk_numeral_i ctx 2)]) (Arithmetic.mk_add ctx [ (List.nth xs 1); (Integer.mk_numeral_i ctx 2)])
(Arithmetic.mk_add ctx [ (List.nth xs 2); (Integer.mk_numeral_i ctx 3)])) ]) in (Arithmetic.mk_add ctx [ (List.nth xs 2); (Integer.mk_numeral_i ctx 3)])) ]) in
let x = (Quantifier.mk_forall ctx types names body_vars (Some 1) [] [] (Some (Symbol.mk_string ctx "Q1")) (Some (Symbol.mk_string ctx "skid1"))) in let x = (Quantifier.mk_forall ctx types names body_vars (Some 1) [] [] (Some (Symbol.mk_string ctx "Q1")) (Some (Symbol.mk_string ctx "skid1"))) in
Printf.printf "Quantifier X: %s\n" (Quantifier.to_string x) ; Printf.printf "Quantifier X: %s\n" (Quantifier.to_string x) ;
let y = (Quantifier.mk_forall_const ctx xs body_const (Some 1) [] [] (Some (Symbol.mk_string ctx "Q2")) (Some (Symbol.mk_string ctx "skid2"))) in let y = (Quantifier.mk_forall_const ctx xs body_const (Some 1) [] [] (Some (Symbol.mk_string ctx "Q2")) (Some (Symbol.mk_string ctx "skid2"))) in
@ -242,8 +241,8 @@ let quantifier_example1 ( ctx : context ) =
open Z3.FloatingPoint open Z3.FloatingPoint
(** (**
A basic example of floating point arithmetic A basic example of floating point arithmetic
**) **)
let fpa_example ( ctx : context ) = let fpa_example ( ctx : context ) =
Printf.printf "FPAExample\n" ; Printf.printf "FPAExample\n" ;
@ -271,7 +270,7 @@ let fpa_example ( ctx : context ) =
(Boolean.mk_not ctx (mk_is_nan ctx y)) ; (Boolean.mk_not ctx (mk_is_nan ctx y)) ;
(Boolean.mk_not ctx (mk_is_infinite ctx y)) ] in (Boolean.mk_not ctx (mk_is_infinite ctx y)) ] in
let args3 = [ c3 ; (Boolean.mk_and ctx and_args) ] in let args3 = [ c3 ; (Boolean.mk_and ctx and_args) ] in
let c4 = (Boolean.mk_and ctx args3) in let c4 = (Boolean.mk_and ctx args3) in
(Printf.printf "c4: %s\n" (Expr.to_string c4)) ; (Printf.printf "c4: %s\n" (Expr.to_string c4)) ;
( (
let solver = (mk_solver ctx None) in let solver = (mk_solver ctx None) in
@ -293,7 +292,7 @@ let fpa_example ( ctx : context ) =
let c2 = (mk_to_fp_bv ctx let c2 = (mk_to_fp_bv ctx
(mk_numeral_string ctx "4619567317775286272" (BitVector.mk_sort ctx 64)) (mk_numeral_string ctx "4619567317775286272" (BitVector.mk_sort ctx 64))
(mk_sort ctx 11 53)) in (mk_sort ctx 11 53)) in
let c3 = (mk_to_fp_int_real ctx let c3 = (mk_to_fp_int_real ctx
(RoundingMode.mk_rtz ctx) (RoundingMode.mk_rtz ctx)
(mk_numeral_string ctx "2" (Integer.mk_sort ctx)) (mk_numeral_string ctx "2" (Integer.mk_sort ctx))
(mk_numeral_string ctx "1.75" (Real.mk_sort ctx)) (mk_numeral_string ctx "1.75" (Real.mk_sort ctx))
@ -304,7 +303,7 @@ let fpa_example ( ctx : context ) =
let args3 = [ (mk_eq ctx c1 c2) ; let args3 = [ (mk_eq ctx c1 c2) ;
(mk_eq ctx c1 c3) ; (mk_eq ctx c1 c3) ;
(mk_eq ctx c1 c4) ] in (mk_eq ctx c1 c4) ] in
let c5 = (Boolean.mk_and ctx args3) in let c5 = (Boolean.mk_and ctx args3) in
(Printf.printf "c5: %s\n" (Expr.to_string c5)) ; (Printf.printf "c5: %s\n" (Expr.to_string c5)) ;
( (
let solver = (mk_solver ctx None) in let solver = (mk_solver ctx None) in
@ -313,7 +312,7 @@ let fpa_example ( ctx : context ) =
raise (TestFailedException "") raise (TestFailedException "")
else else
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
) )
(** (**
A basic example of RCF usage A basic example of RCF usage
@ -322,11 +321,58 @@ let rcf_example ( ctx : context ) =
Printf.printf "RCFExample\n" ; Printf.printf "RCFExample\n" ;
let pi = RCF.mk_pi ctx in let pi = RCF.mk_pi ctx in
let e = RCF.mk_e ctx in let e = RCF.mk_e ctx in
let inf0 = RCF.mk_infinitesimal ctx in
let inf1 = RCF.mk_infinitesimal ctx in
let r = RCF.mk_rational ctx "42.001" in
let pi_div_e = RCF.div ctx pi e in
let pi_div_r = RCF.div ctx pi r in
(Printf.printf "e: %s, pi: %s, e==pi: %b, e < pi: %b\n" (Printf.printf "e: %s, pi: %s, e==pi: %b, e < pi: %b\n"
(RCF.num_to_string ctx e true false) (RCF.num_to_string ctx e true false)
(RCF.num_to_string ctx pi true false) (RCF.num_to_string ctx pi true false)
(RCF.eq ctx e pi) (RCF.eq ctx e pi)
(RCF.lt ctx e pi)) ; (RCF.lt ctx e pi)) ;
Printf.printf "pi_div_e: %s.\n" (RCF.num_to_string ctx pi_div_e true false);
Printf.printf "pi_div_r: %s.\n" (RCF.num_to_string ctx pi_div_r true false);
Printf.printf "inf0: %s.\n" (RCF.num_to_string ctx inf0 true false);
Printf.printf "(RCF.is_rational ctx pi): %b.\n" (RCF.is_rational ctx pi);
Printf.printf "(RCF.is_algebraic ctx pi): %b.\n" (RCF.is_algebraic ctx pi);
Printf.printf "(RCF.is_transcendental ctx pi): %b.\n" (RCF.is_transcendental ctx pi);
Printf.printf "(RCF.is_rational ctx r): %b.\n" (RCF.is_rational ctx r);
Printf.printf "(RCF.is_algebraic ctx r): %b.\n" (RCF.is_algebraic ctx r);
Printf.printf "(RCF.is_transcendental ctx r): %b.\n" (RCF.is_transcendental ctx r);
Printf.printf "(RCF.is_infinitesimal ctx inf0): %b.\n" (RCF.is_infinitesimal ctx inf0);
Printf.printf "(RCF.extension_index ctx inf0): %d.\n" (RCF.extension_index ctx inf0);
Printf.printf "(RCF.extension_index ctx inf1): %d.\n" (RCF.extension_index ctx inf1);
let poly:RCF.rcf_num list = [ e; pi; inf0 ] in
let rs:RCF.root list = RCF.roots ctx poly in
let print_root (x:RCF.root) =
begin
Printf.printf "root: %s\n%!" (RCF.num_to_string ctx x.obj true false);
if RCF.is_algebraic ctx x.obj then (
(match x.interval with
| Some ivl -> Printf.printf " interval: (%b, %b, %s, %b, %b, %s)\n"
ivl.lower_is_inf
ivl.lower_is_open
(RCF.num_to_string ctx ivl.lower true false)
ivl.upper_is_inf
ivl.upper_is_open
(RCF.num_to_string ctx ivl.upper true false);
| None -> ());
Printf.printf " polynomial coefficients:";
List.iter (fun c -> Printf.printf " %s" (RCF.num_to_string ctx c false false)) x.polynomial;
Printf.printf "\n";
Printf.printf " sign conditions:";
List.iter
(fun (poly, sign) ->
List.iter (fun p -> Printf.printf " %s" (RCF.num_to_string ctx p true false)) poly;
Printf.printf " %s" (if sign > 0 then "> 0" else if sign < 0 then "< 0" else "= 0"))
x.sign_conditions;
Printf.printf "\n")
end
in
List.iter print_root rs;
RCF.del_roots ctx rs;
RCF.del_list ctx [pi; e; inf0; inf1; r; pi_div_e; pi_div_r];
Printf.printf "Test passed.\n" Printf.printf "Test passed.\n"
let _ = let _ =
@ -363,5 +409,6 @@ let _ =
) with Error(msg) -> ( ) with Error(msg) -> (
Printf.printf "Z3 EXCEPTION: %s\n" msg ; Printf.printf "Z3 EXCEPTION: %s\n" msg ;
exit 1 exit 1
) )
;; ;;

View file

@ -17,7 +17,7 @@ Author:
Leonardo de Moura (leonardo) 2012-01-05 Leonardo de Moura (leonardo) 2012-01-05
Notes: Notes:
--*/ --*/
#include "api/z3.h" #include "api/z3.h"
#include "api/api_log_macros.h" #include "api/api_log_macros.h"
@ -32,12 +32,12 @@ static void reset_rcf_cancel(Z3_context c) {
// no-op // no-op
} }
static Z3_rcf_num from_rcnumeral(rcnumeral a) { static Z3_rcf_num from_rcnumeral(rcnumeral a) {
return reinterpret_cast<Z3_rcf_num>(a.data()); return reinterpret_cast<Z3_rcf_num>(a.data());
} }
static rcnumeral to_rcnumeral(Z3_rcf_num a) { static rcnumeral to_rcnumeral(Z3_rcf_num a) {
return rcnumeral::mk(a); return rcnumeral::mk(a);
} }
extern "C" { extern "C" {
@ -179,7 +179,7 @@ extern "C" {
RETURN_Z3(from_rcnumeral(r)); RETURN_Z3(from_rcnumeral(r));
Z3_CATCH_RETURN(nullptr); Z3_CATCH_RETURN(nullptr);
} }
Z3_rcf_num Z3_API Z3_rcf_neg(Z3_context c, Z3_rcf_num a) { Z3_rcf_num Z3_API Z3_rcf_neg(Z3_context c, Z3_rcf_num a) {
Z3_TRY; Z3_TRY;
LOG_Z3_rcf_neg(c, a); LOG_Z3_rcf_neg(c, a);
@ -302,4 +302,139 @@ extern "C" {
Z3_CATCH; Z3_CATCH;
} }
bool Z3_API Z3_rcf_is_rational(Z3_context c, Z3_rcf_num a) {
Z3_TRY;
LOG_Z3_rcf_is_rational(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).is_rational(to_rcnumeral(a));
Z3_CATCH_RETURN(false);
}
bool Z3_API Z3_rcf_is_algebraic(Z3_context c, Z3_rcf_num a) {
Z3_TRY;
LOG_Z3_rcf_is_algebraic(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).is_algebraic(to_rcnumeral(a));
Z3_CATCH_RETURN(false);
}
bool Z3_API Z3_rcf_is_infinitesimal(Z3_context c, Z3_rcf_num a) {
Z3_TRY;
LOG_Z3_rcf_is_infinitesimal(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).is_infinitesimal(to_rcnumeral(a));
Z3_CATCH_RETURN(false);
}
bool Z3_API Z3_rcf_is_transcendental(Z3_context c, Z3_rcf_num a) {
Z3_TRY;
LOG_Z3_rcf_is_transcendental(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).is_transcendental(to_rcnumeral(a));
Z3_CATCH_RETURN(false);
}
unsigned Z3_API Z3_rcf_extension_index(Z3_context c, Z3_rcf_num a) {
Z3_TRY;
LOG_Z3_rcf_extension_index(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).extension_index(to_rcnumeral(a));
Z3_CATCH_RETURN(false);
}
Z3_symbol Z3_API Z3_rcf_transcendental_name(Z3_context c, Z3_rcf_num a) {
Z3_TRY;
LOG_Z3_rcf_transcendental_name(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return of_symbol(rcfm(c).transcendental_name(to_rcnumeral(a)));
Z3_CATCH_RETURN(of_symbol(symbol::null));
}
Z3_symbol Z3_API Z3_rcf_infinitesimal_name(Z3_context c, Z3_rcf_num a) {
Z3_TRY;
LOG_Z3_rcf_infinitesimal_name(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return of_symbol(rcfm(c).infinitesimal_name(to_rcnumeral(a)));
Z3_CATCH_RETURN(of_symbol(symbol::null));
}
unsigned Z3_API Z3_rcf_num_coefficients(Z3_context c, Z3_rcf_num a)
{
Z3_TRY;
LOG_Z3_rcf_num_coefficients(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).num_coefficients(to_rcnumeral(a));
Z3_CATCH_RETURN(0);
}
Z3_rcf_num Z3_API Z3_rcf_coefficient(Z3_context c, Z3_rcf_num a, unsigned i)
{
Z3_TRY;
LOG_Z3_rcf_coefficient(c, a, i);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return from_rcnumeral(rcfm(c).get_coefficient(to_rcnumeral(a), i));
Z3_CATCH_RETURN(nullptr);
}
int Z3_API Z3_rcf_interval(Z3_context c, Z3_rcf_num a, int * lower_is_inf, int * lower_is_open, Z3_rcf_num * lower, int * upper_is_inf, int * upper_is_open, Z3_rcf_num * upper) {
Z3_TRY;
LOG_Z3_rcf_interval(c, a, lower_is_inf, lower_is_open, lower, upper_is_inf, upper_is_open, upper);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
rcnumeral r_lower, r_upper;
bool r = rcfm(c).get_interval(to_rcnumeral(a), *lower_is_inf, *lower_is_open, r_lower, *upper_is_inf, *upper_is_open, r_upper);
*lower = from_rcnumeral(r_lower);
*upper = from_rcnumeral(r_upper);
return r;
Z3_CATCH_RETURN(0);
}
unsigned Z3_API Z3_rcf_num_sign_conditions(Z3_context c, Z3_rcf_num a)
{
Z3_TRY;
LOG_Z3_rcf_num_sign_conditions(c, a);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).num_sign_conditions(to_rcnumeral(a));
Z3_CATCH_RETURN(0);
}
int Z3_API Z3_rcf_sign_condition_sign(Z3_context c, Z3_rcf_num a, unsigned i)
{
Z3_TRY;
LOG_Z3_rcf_sign_condition_sign(c, a, i);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).get_sign_condition_sign(to_rcnumeral(a), i);
Z3_CATCH_RETURN(0);
}
unsigned Z3_API Z3_rcf_num_sign_condition_coefficients(Z3_context c, Z3_rcf_num a, unsigned i)
{
Z3_TRY;
LOG_Z3_rcf_num_sign_condition_coefficients(c, a, i);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return rcfm(c).num_sign_condition_coefficients(to_rcnumeral(a), i);
Z3_CATCH_RETURN(0);
}
Z3_rcf_num Z3_API Z3_rcf_sign_condition_coefficient(Z3_context c, Z3_rcf_num a, unsigned i, unsigned j)
{
Z3_TRY;
LOG_Z3_rcf_sign_condition_coefficient(c, a, i, j);
RESET_ERROR_CODE();
reset_rcf_cancel(c);
return from_rcnumeral(rcfm(c).get_sign_condition_coefficient(to_rcnumeral(a), i, j));
Z3_CATCH_RETURN(nullptr);
}
}; };

View file

@ -2080,6 +2080,7 @@ struct
type rcf_num = Z3native.rcf_num type rcf_num = Z3native.rcf_num
let del (ctx:context) (a:rcf_num) = Z3native.rcf_del ctx a let del (ctx:context) (a:rcf_num) = Z3native.rcf_del ctx a
let del_list (ctx:context) (ns:rcf_num list) = List.iter (fun a -> Z3native.rcf_del ctx a) ns
let mk_rational (ctx:context) (v:string) = Z3native.rcf_mk_rational ctx v let mk_rational (ctx:context) (v:string) = Z3native.rcf_mk_rational ctx v
let mk_small_int (ctx:context) (v:int) = Z3native.rcf_mk_small_int ctx v let mk_small_int (ctx:context) (v:int) = Z3native.rcf_mk_small_int ctx v
@ -2087,7 +2088,9 @@ struct
let mk_e (ctx:context) = Z3native.rcf_mk_e ctx let mk_e (ctx:context) = Z3native.rcf_mk_e ctx
let mk_infinitesimal (ctx:context) = Z3native.rcf_mk_infinitesimal ctx let mk_infinitesimal (ctx:context) = Z3native.rcf_mk_infinitesimal ctx
let mk_roots (ctx:context) (n:int) (a:rcf_num list) = let n, r = Z3native.rcf_mk_roots ctx n a in r let mk_roots (ctx:context) (a:rcf_num list) =
let n, r = Z3native.rcf_mk_roots ctx (List.length a) a in
List.init n (fun x -> List.nth r x)
let add (ctx:context) (a:rcf_num) (b:rcf_num) = Z3native.rcf_add ctx a b let add (ctx:context) (a:rcf_num) (b:rcf_num) = Z3native.rcf_add ctx a b
let sub (ctx:context) (a:rcf_num) (b:rcf_num) = Z3native.rcf_sub ctx a b let sub (ctx:context) (a:rcf_num) (b:rcf_num) = Z3native.rcf_sub ctx a b
@ -2109,6 +2112,83 @@ struct
let num_to_string (ctx:context) (a:rcf_num) (compact:bool) (html:bool) = Z3native.rcf_num_to_string ctx a compact html let num_to_string (ctx:context) (a:rcf_num) (compact:bool) (html:bool) = Z3native.rcf_num_to_string ctx a compact html
let num_to_decimal_string (ctx:context) (a:rcf_num) (prec:int) = Z3native.rcf_num_to_decimal_string ctx a prec let num_to_decimal_string (ctx:context) (a:rcf_num) (prec:int) = Z3native.rcf_num_to_decimal_string ctx a prec
let get_numerator_denominator (ctx:context) (a:rcf_num) = Z3native.rcf_get_numerator_denominator ctx a let get_numerator_denominator (ctx:context) (a:rcf_num) = Z3native.rcf_get_numerator_denominator ctx a
let is_rational (ctx:context) (a:rcf_num) = Z3native.rcf_is_rational ctx a
let is_algebraic (ctx:context) (a:rcf_num) = Z3native.rcf_is_algebraic ctx a
let is_infinitesimal (ctx:context) (a:rcf_num) = Z3native.rcf_is_infinitesimal ctx a
let is_transcendental (ctx:context) (a:rcf_num) = Z3native.rcf_is_transcendental ctx a
let extension_index (ctx:context) (a:rcf_num) = Z3native.rcf_extension_index ctx a
let transcendental_name (ctx:context) (a:rcf_num) = Z3native.rcf_transcendental_name ctx a
let infinitesimal_name (ctx:context) (a:rcf_num) = Z3native.rcf_infinitesimal_name ctx a
let num_coefficients (ctx:context) (a:rcf_num) = Z3native.rcf_num_coefficients ctx a
let get_coefficient (ctx:context) (a:rcf_num) (i:int) = Z3native.rcf_coefficient ctx a i
let coefficients (ctx:context) (a:rcf_num) =
List.init (num_coefficients ctx a) (fun i -> Z3native.rcf_coefficient ctx a i)
type interval = {
lower_is_inf : bool;
lower_is_open : bool;
lower : rcf_num;
upper_is_inf : bool;
upper_is_open : bool;
upper : rcf_num;
}
let root_interval (ctx:context) (a:rcf_num) =
let ok, linf, lopen, l, uinf, uopen, u = Z3native.rcf_interval ctx a in
let i:interval = {
lower_is_inf = linf != 0;
lower_is_open = lopen != 0;
lower = l;
upper_is_inf = uinf != 0;
upper_is_open = uopen != 0;
upper = u } in
if ok != 0 then Some i else None
let sign_condition_sign (ctx:context) (a:rcf_num) (i:int) = Z3native.rcf_sign_condition_sign ctx a i
let sign_condition_coefficient (ctx:context) (a:rcf_num) (i:int) (j:int) = Z3native.rcf_sign_condition_coefficient ctx a i j
let num_sign_condition_coefficients (ctx:context) (a:rcf_num) (i:int) = Z3native.rcf_num_sign_condition_coefficients ctx a i
let sign_condition_coefficients (ctx:context) (a:rcf_num) (i:int) =
let n = Z3native.rcf_num_sign_condition_coefficients ctx a i in
List.init n (fun j -> Z3native.rcf_sign_condition_coefficient ctx a i j)
let sign_conditions (ctx:context) (a:rcf_num) =
let n = Z3native.rcf_num_sign_conditions ctx a in
List.init n (fun i ->
(let nc = Z3native.rcf_num_sign_condition_coefficients ctx a i in
List.init nc (fun j -> Z3native.rcf_sign_condition_coefficient ctx a i j)),
Z3native.rcf_sign_condition_sign ctx a i)
type root = {
obj : rcf_num;
polynomial : rcf_num list;
interval : interval option;
sign_conditions : (rcf_num list * int) list;
}
let roots (ctx:context) (a:rcf_num list) =
let rs = mk_roots ctx a in
List.map
(fun r -> {
obj = r;
polynomial = coefficients ctx r;
interval = root_interval ctx r;
sign_conditions = sign_conditions ctx r})
rs
let del_root (ctx:context) (r:root) =
del ctx r.obj;
List.iter (fun n -> del ctx n) r.polynomial;
List.iter (fun (ns, _) -> del_list ctx ns) r.sign_conditions
let del_roots (ctx:context) (rs:root list) =
List.iter (fun r -> del_root ctx r) rs
end end

View file

@ -3553,77 +3553,147 @@ end
(** Real closed field *) (** Real closed field *)
module RCF : module RCF :
sig sig
type rcf_num type rcf_num
(** Delete a RCF numeral created using the RCF API. *) (** Delete a RCF numeral created using the RCF API. *)
val del : context -> rcf_num -> unit val del : context -> rcf_num -> unit
(** Return a RCF rational using the given string. *) (** Delete RCF numerals created using the RCF API. *)
val mk_rational : context -> string -> rcf_num val del_list : context -> rcf_num list -> unit
(** Return a RCF small integer. *) (** Return a RCF rational using the given string. *)
val mk_small_int : context -> int -> rcf_num val mk_rational : context -> string -> rcf_num
(** Return Pi *) (** Return a RCF small integer. *)
val mk_pi : context -> rcf_num val mk_small_int : context -> int -> rcf_num
(** Return e (Euler's constant) *) (** Return Pi *)
val mk_e : context -> rcf_num val mk_pi : context -> rcf_num
(** Return a new infinitesimal that is smaller than all elements in the Z3 field. *) (** Return e (Euler's constant) *)
val mk_infinitesimal : context -> rcf_num val mk_e : context -> rcf_num
(** Extract the roots of a polynomial. Precondition: The input polynomial is not the zero polynomial. *) (** Return a new infinitesimal that is smaller than all elements in the Z3 field. *)
val mk_roots : context -> int -> rcf_num list -> rcf_num list val mk_infinitesimal : context -> rcf_num
(** Addition *) (** Extract the roots of a polynomial. Precondition: The input polynomial is not the zero polynomial. *)
val add : context -> rcf_num -> rcf_num -> rcf_num val mk_roots : context -> rcf_num list -> rcf_num list
(** Subtraction *) (** Addition *)
val sub : context -> rcf_num -> rcf_num -> rcf_num val add : context -> rcf_num -> rcf_num -> rcf_num
(** Multiplication *) (** Subtraction *)
val mul : context -> rcf_num -> rcf_num -> rcf_num val sub : context -> rcf_num -> rcf_num -> rcf_num
(** Division *) (** Multiplication *)
val div : context -> rcf_num -> rcf_num -> rcf_num val mul : context -> rcf_num -> rcf_num -> rcf_num
(** Negation *) (** Division *)
val neg : context -> rcf_num -> rcf_num val div : context -> rcf_num -> rcf_num -> rcf_num
(** Multiplicative Inverse *) (** Negation *)
val inv : context -> rcf_num -> rcf_num val neg : context -> rcf_num -> rcf_num
(** Power *) (** Multiplicative Inverse *)
val power : context -> rcf_num -> int -> rcf_num val inv : context -> rcf_num -> rcf_num
(** less-than *) (** Power *)
val lt : context -> rcf_num -> rcf_num -> bool val power : context -> rcf_num -> int -> rcf_num
(** greater-than *) (** less-than *)
val gt : context -> rcf_num -> rcf_num -> bool val lt : context -> rcf_num -> rcf_num -> bool
(** less-than or equal *) (** greater-than *)
val le : context -> rcf_num -> rcf_num -> bool val gt : context -> rcf_num -> rcf_num -> bool
(** greater-than or equal *) (** less-than or equal *)
val ge : context -> rcf_num -> rcf_num -> bool val le : context -> rcf_num -> rcf_num -> bool
(** equality *) (** greater-than or equal *)
val eq : context -> rcf_num -> rcf_num -> bool val ge : context -> rcf_num -> rcf_num -> bool
(** not equal *) (** equality *)
val neq : context -> rcf_num -> rcf_num -> bool val eq : context -> rcf_num -> rcf_num -> bool
(** Convert the RCF numeral into a string. *) (** not equal *)
val num_to_string : context -> rcf_num -> bool -> bool -> string val neq : context -> rcf_num -> rcf_num -> bool
(** Convert the RCF numeral into a string in decimal notation. *) (** Convert the RCF numeral into a string. *)
val num_to_decimal_string : context -> rcf_num -> int -> string val num_to_string : context -> rcf_num -> bool -> bool -> string
(** Extract the "numerator" and "denominator" of the given RCF numeral. (** Convert the RCF numeral into a string in decimal notation. *)
We have that \ccode{a = n/d}, moreover \c n and \c d are not represented using rational functions. *) val num_to_decimal_string : context -> rcf_num -> int -> string
val get_numerator_denominator : context -> rcf_num -> (rcf_num * rcf_num)
(** Extract the "numerator" and "denominator" of the given RCF numeral.
We have that \ccode{a = n/d}, moreover \c n and \c d are not represented using rational functions. *)
val get_numerator_denominator : context -> rcf_num -> (rcf_num * rcf_num)
(** Return \c true if \c a represents a rational number. *)
val is_rational : context -> rcf_num -> bool
(** Return \c true if \c a represents an algebraic number. *)
val is_algebraic : context -> rcf_num -> bool
(** Return \c true if \c a represents an infinitesimal. *)
val is_infinitesimal : context -> rcf_num -> bool
(** Return \c true if \c a represents a transcendental number. *)
val is_transcendental : context -> rcf_num -> bool
(** Return the index of a field extension. *)
val extension_index : context -> rcf_num -> int
(** Return the name of a transcendental. *)
val transcendental_name : context -> rcf_num -> Symbol.symbol
(** Return the name of an infinitesimal. *)
val infinitesimal_name : context -> rcf_num -> Symbol.symbol
(** Return the number of coefficients in an algebraic number. *)
val num_coefficients : context -> rcf_num -> int
(** Extract a coefficient from an algebraic number. *)
val get_coefficient : context -> rcf_num -> int -> rcf_num
(** Extract the coefficients from an algebraic number. *)
val coefficients : context -> rcf_num -> rcf_num list
(** Extract the sign of a sign condition from an algebraic number. *)
val sign_condition_sign : context -> rcf_num -> int -> int
(** Return the size of a sign condition polynomial. *)
val num_sign_condition_coefficients : context -> rcf_num -> int -> int
(** Extract a sign condition polynomial coefficient from an algebraic number. *)
val sign_condition_coefficient : context -> rcf_num -> int -> int -> rcf_num
(** Extract sign conditions from an algebraic number. *)
val sign_conditions : context -> rcf_num -> (rcf_num list * int) list
(** Extract the interval from an algebraic number. *)
type interval = {
lower_is_inf : bool;
lower_is_open : bool;
lower : rcf_num;
upper_is_inf : bool;
upper_is_open : bool;
upper : rcf_num;
}
val root_interval : context -> rcf_num -> interval option
type root = {
obj : rcf_num;
polynomial : rcf_num list;
interval : interval option;
sign_conditions : (rcf_num list * int) list;
}
val roots : context -> rcf_num list -> root list
val del_root : context -> root -> unit
val del_roots : context -> root list -> unit
end end
(** Set a global (or module) parameter, which is shared by all Z3 contexts. (** Set a global (or module) parameter, which is shared by all Z3 contexts.

View file

@ -193,8 +193,124 @@ extern "C" {
We have that \ccode{a = n/d}, moreover \c n and \c d are not represented using rational functions. We have that \ccode{a = n/d}, moreover \c n and \c d are not represented using rational functions.
def_API('Z3_rcf_get_numerator_denominator', VOID, (_in(CONTEXT), _in(RCF_NUM), _out(RCF_NUM), _out(RCF_NUM))) def_API('Z3_rcf_get_numerator_denominator', VOID, (_in(CONTEXT), _in(RCF_NUM), _out(RCF_NUM), _out(RCF_NUM)))
*/ */
void Z3_API Z3_rcf_get_numerator_denominator(Z3_context c, Z3_rcf_num a, Z3_rcf_num * n, Z3_rcf_num * d); void Z3_API Z3_rcf_get_numerator_denominator(Z3_context c, Z3_rcf_num a, Z3_rcf_num * n, Z3_rcf_num * d);
/**
\brief Return \c true if \c a represents a rational number.
def_API('Z3_rcf_is_rational', BOOL, (_in(CONTEXT), _in(RCF_NUM)))
*/
bool Z3_API Z3_rcf_is_rational(Z3_context c, Z3_rcf_num a);
/**
\brief Return \c true if \c a represents an algebraic number.
def_API('Z3_rcf_is_algebraic', BOOL, (_in(CONTEXT), _in(RCF_NUM)))
*/
bool Z3_API Z3_rcf_is_algebraic(Z3_context c, Z3_rcf_num a);
/**
\brief Return \c true if \c a represents an infinitesimal.
def_API('Z3_rcf_is_infinitesimal', BOOL, (_in(CONTEXT), _in(RCF_NUM)))
*/
bool Z3_API Z3_rcf_is_infinitesimal(Z3_context c, Z3_rcf_num a);
/**
\brief Return \c true if \c a represents a transcendental number.
def_API('Z3_rcf_is_transcendental', BOOL, (_in(CONTEXT), _in(RCF_NUM)))
*/
bool Z3_API Z3_rcf_is_transcendental(Z3_context c, Z3_rcf_num a);
/**
\brief Return the index of a field extension.
def_API('Z3_rcf_extension_index', UINT, (_in(CONTEXT), _in(RCF_NUM)))
*/
unsigned Z3_API Z3_rcf_extension_index(Z3_context c, Z3_rcf_num a);
/**
\brief Return the name of a transcendental.
\pre Z3_rcf_is_transcendtal(ctx, a);
def_API('Z3_rcf_transcendental_name', SYMBOL, (_in(CONTEXT), _in(RCF_NUM)))
*/
Z3_symbol Z3_API Z3_rcf_transcendental_name(Z3_context c, Z3_rcf_num a);
/**
\brief Return the name of an infinitesimal.
\pre Z3_rcf_is_infinitesimal(ctx, a);
def_API('Z3_rcf_infinitesimal_name', SYMBOL, (_in(CONTEXT), _in(RCF_NUM)))
*/
Z3_symbol Z3_API Z3_rcf_infinitesimal_name(Z3_context c, Z3_rcf_num a);
/**
\brief Return the number of coefficients in an algebraic number.
\pre Z3_rcf_is_algebraic(ctx, a);
def_API('Z3_rcf_num_coefficients', UINT, (_in(CONTEXT), _in(RCF_NUM)))
*/
unsigned Z3_API Z3_rcf_num_coefficients(Z3_context c, Z3_rcf_num a);
/**
\brief Extract a coefficient from an algebraic number.
\pre Z3_rcf_is_algebraic(ctx, a);
def_API('Z3_rcf_coefficient', RCF_NUM, (_in(CONTEXT), _in(RCF_NUM), _in(UINT)))
*/
Z3_rcf_num Z3_API Z3_rcf_coefficient(Z3_context c, Z3_rcf_num a, unsigned i);
/**
\brief Extract an interval from an algebraic number.
\pre Z3_rcf_is_algebraic(ctx, a);
def_API('Z3_rcf_interval', INT, (_in(CONTEXT), _in(RCF_NUM), _out(INT), _out(INT), _out(RCF_NUM), _out(INT), _out(INT), _out(RCF_NUM)))
*/
int Z3_API Z3_rcf_interval(Z3_context c, Z3_rcf_num a, int * lower_is_inf, int * lower_is_open, Z3_rcf_num * lower, int * upper_is_inf, int * upper_is_open, Z3_rcf_num * upper);
/**
\brief Return the number of sign conditions of an algebraic number.
\pre Z3_rcf_is_algebraic(ctx, a);
def_API('Z3_rcf_num_sign_conditions', UINT, (_in(CONTEXT), _in(RCF_NUM)))
*/
unsigned Z3_API Z3_rcf_num_sign_conditions(Z3_context c, Z3_rcf_num a);
/**
\brief Extract the sign of a sign condition from an algebraic number.
\pre Z3_rcf_is_algebraic(ctx, a);
def_API('Z3_rcf_sign_condition_sign', INT, (_in(CONTEXT), _in(RCF_NUM), _in(UINT)))
*/
int Z3_API Z3_rcf_sign_condition_sign(Z3_context c, Z3_rcf_num a, unsigned i);
/**
\brief Return the number of sign condition polynomial coefficients of an algebraic number.
\pre Z3_rcf_is_algebraic(ctx, a);
def_API('Z3_rcf_num_sign_condition_coefficients', UINT, (_in(CONTEXT), _in(RCF_NUM), _in(UINT)))
*/
unsigned Z3_API Z3_rcf_num_sign_condition_coefficients(Z3_context c, Z3_rcf_num a, unsigned i);
/**
\brief Extract the j-th polynomial coefficient of the i-th sign condition.
\pre Z3_rcf_is_algebraic(ctx, a);
def_API('Z3_rcf_sign_condition_coefficient', RCF_NUM, (_in(CONTEXT), _in(RCF_NUM), _in(UINT), _in(UINT)))
*/
Z3_rcf_num Z3_API Z3_rcf_sign_condition_coefficient(Z3_context c, Z3_rcf_num a, unsigned i, unsigned j);
/**@}*/ /**@}*/
/**@}*/ /**@}*/

View file

@ -2498,6 +2498,35 @@ namespace realclosure {
} }
} }
/**
\brief Return true if a is an algebraic number.
*/
bool is_algebraic(numeral const & a) {
return is_rational_function(a) && to_rational_function(a)->ext()->is_algebraic();
}
/**
\brief Return true if a represents an infinitesimal.
*/
bool is_infinitesimal(numeral const & a) {
return is_rational_function(a) && to_rational_function(a)->ext()->is_infinitesimal();
}
/**
\brief Return true if a is a transcendental.
*/
bool is_transcendental(numeral const & a) {
return is_rational_function(a) && to_rational_function(a)->ext()->is_transcendental();
}
/**
\brief Return true if a is a rational.
*/
bool is_rational(numeral const & a) {
return a.m_value->is_rational();
}
/** /**
\brief Return true if a depends on infinitesimal extensions. \brief Return true if a depends on infinitesimal extensions.
*/ */
@ -3330,6 +3359,151 @@ namespace realclosure {
set(q, _q); set(q, _q);
} }
unsigned extension_index(numeral const & a) {
if (!is_rational_function(a))
return -1;
return to_rational_function(a)->ext()->idx();
}
symbol transcendental_name(numeral const & a) {
if (!is_transcendental(a))
return symbol();
return to_transcendental(to_rational_function(a)->ext())->m_name;
}
symbol infinitesimal_name(numeral const & a) {
if (!is_infinitesimal(a))
return symbol();
return to_infinitesimal(to_rational_function(a)->ext())->m_name;
}
unsigned num_coefficients(numeral const & a) {
if (!is_algebraic(a))
return 0;
return to_algebraic(to_rational_function(a)->ext())->p().size();
}
numeral get_coefficient(numeral const & a, unsigned i)
{
if (!is_algebraic(a))
return numeral();
algebraic * ext = to_algebraic(to_rational_function(a)->ext());
if (i >= ext->p().size())
return numeral();
value_ref v(*this);
v = ext->p()[i];
numeral r;
set(r, v);
return r;
}
unsigned num_sign_conditions(numeral const & a) {
unsigned r = 0;
if (is_algebraic(a)) {
algebraic * ext = to_algebraic(to_rational_function(a)->ext());
const sign_det * sdt = ext->sdt();
if (sdt) {
sign_condition * sc = sdt->sc(ext->sc_idx());
while (sc) {
r++;
sc = sc->prev();
}
}
}
return r;
}
int get_sign_condition_sign(numeral const & a, unsigned i)
{
if (!is_algebraic(a))
return 0;
algebraic * ext = to_algebraic(to_rational_function(a)->ext());
const sign_det * sdt = ext->sdt();
if (!sdt)
return 0;
else {
sign_condition * sc = sdt->sc(ext->sc_idx());
while (i) {
if (sc) sc = sc->prev();
i--;
}
return sc ? sc->sign() : 0;
}
}
bool get_interval(numeral const & a, int & lower_is_inf, int & lower_is_open, numeral & lower, int & upper_is_inf, int & upper_is_open, numeral & upper)
{
if (!is_algebraic(a))
return false;
lower = numeral();
upper = numeral();
algebraic * ext = to_algebraic(to_rational_function(a)->ext());
mpbqi &ivl = ext->iso_interval();
lower_is_inf = ivl.lower_is_inf();
lower_is_open = ivl.lower_is_open();
if (!m_bqm.is_zero(ivl.lower()))
set(lower, mk_rational(ivl.lower()));
upper_is_inf = ivl.upper_is_inf();
upper_is_open = ivl.upper_is_open();
if (!m_bqm.is_zero(ivl.upper()))
set(upper, mk_rational(ivl.upper()));
return true;
}
unsigned get_sign_condition_size(numeral const &a, unsigned i) {
algebraic * ext = to_algebraic(to_rational_function(a)->ext());
const sign_det * sdt = ext->sdt();
if (!sdt)
return 0;
sign_condition * sc = sdt->sc(ext->sc_idx());
while (i) {
if (sc) sc = sc->prev();
i--;
}
return ext->sdt()->qs()[sc->qidx()].size();
}
int num_sign_condition_coefficients(numeral const &a, unsigned i)
{
if (!is_algebraic(a))
return 0;
algebraic * ext = to_algebraic(to_rational_function(a)->ext());
const sign_det * sdt = ext->sdt();
if (!sdt)
return 0;
sign_condition * sc = sdt->sc(ext->sc_idx());
while (i) {
if (sc) sc = sc->prev();
i--;
}
const polynomial & q = ext->sdt()->qs()[sc->qidx()];
return q.size();
}
numeral get_sign_condition_coefficient(numeral const &a, unsigned i, unsigned j)
{
if (!is_algebraic(a))
return numeral();
algebraic * ext = to_algebraic(to_rational_function(a)->ext());
const sign_det * sdt = ext->sdt();
if (!sdt)
return numeral();
sign_condition * sc = sdt->sc(ext->sc_idx());
while (i) {
if (sc) sc = sc->prev();
i--;
}
const polynomial & q = ext->sdt()->qs()[sc->qidx()];
if (j >= q.size())
return numeral();
value_ref v(*this);
v = q[j];
numeral r;
set(r, v);
return r;
}
// --------------------------------- // ---------------------------------
// //
// GCD of integer coefficients // GCD of integer coefficients
@ -6103,6 +6277,22 @@ namespace realclosure {
return m_imp->is_int(a); return m_imp->is_int(a);
} }
bool manager::is_rational(numeral const & a) {
return m_imp->is_rational(a);
}
bool manager::is_algebraic(numeral const & a) {
return m_imp->is_algebraic(a);
}
bool manager::is_infinitesimal(numeral const & a) {
return m_imp->is_infinitesimal(a);
}
bool manager::is_transcendental(numeral const & a) {
return m_imp->is_transcendental(a);
}
bool manager::depends_on_infinitesimals(numeral const & a) { bool manager::depends_on_infinitesimals(numeral const & a) {
return m_imp->depends_on_infinitesimals(a); return m_imp->depends_on_infinitesimals(a);
} }
@ -6251,6 +6441,56 @@ namespace realclosure {
save_interval_ctx ctx(this); save_interval_ctx ctx(this);
m_imp->clean_denominators(a, p, q); m_imp->clean_denominators(a, p, q);
} }
unsigned manager::extension_index(numeral const & a)
{
return m_imp->extension_index(a);
}
symbol manager::transcendental_name(numeral const &a)
{
return m_imp->transcendental_name(a);
}
symbol manager::infinitesimal_name(numeral const &a)
{
return m_imp->infinitesimal_name(a);
}
unsigned manager::num_coefficients(numeral const &a)
{
return m_imp->num_coefficients(a);
}
manager::numeral manager::get_coefficient(numeral const &a, unsigned i)
{
return m_imp->get_coefficient(a, i);
}
unsigned manager::num_sign_conditions(numeral const &a)
{
return m_imp->num_sign_conditions(a);
}
int manager::get_sign_condition_sign(numeral const &a, unsigned i)
{
return m_imp->get_sign_condition_sign(a, i);
}
bool manager::get_interval(numeral const & a, int & lower_is_inf, int & lower_is_open, numeral & lower, int & upper_is_inf, int & upper_is_open, numeral & upper)
{
return m_imp->get_interval(a, lower_is_inf, lower_is_open, lower, upper_is_inf, upper_is_open, upper);
}
unsigned manager::num_sign_condition_coefficients(numeral const &a, unsigned i)
{
return m_imp->num_sign_condition_coefficients(a, i);
}
manager::numeral manager::get_sign_condition_coefficient(numeral const &a, unsigned i, unsigned j)
{
return m_imp->get_sign_condition_coefficient(a, i, j);
}
}; };
void pp(realclosure::manager::imp * imp, realclosure::polynomial const & p, realclosure::extension * ext) { void pp(realclosure::manager::imp * imp, realclosure::polynomial const & p, realclosure::extension * ext) {

View file

@ -70,14 +70,14 @@ namespace realclosure {
*/ */
void mk_infinitesimal(char const * name, char const * pp_name, numeral & r); void mk_infinitesimal(char const * name, char const * pp_name, numeral & r);
void mk_infinitesimal(numeral & r); void mk_infinitesimal(numeral & r);
/** /**
\brief Add a new transcendental real value to the field. \brief Add a new transcendental real value to the field.
The functor \c mk_interval is used to compute approximations of the transcendental value. The functor \c mk_interval is used to compute approximations of the transcendental value.
This procedure should be used with care, if the value is not really transcendental with respect to the current This procedure should be used with care, if the value is not really transcendental with respect to the current
field, computations with the new numeral may not terminate. field, computations with the new numeral may not terminate.
Example: we extended the field with Pi. Pi is transcendental with respect to a field that contains only algebraic real numbers. Example: we extended the field with Pi. Pi is transcendental with respect to a field that contains only algebraic real numbers.
So, this step is fine. Let us call the resultant field F. So, this step is fine. Let us call the resultant field F.
Then, we extend the field F with 1 - Pi. 1 - Pi is transcendental with respect to algebraic real numbers, but it is NOT transcendental Then, we extend the field F with 1 - Pi. 1 - Pi is transcendental with respect to algebraic real numbers, but it is NOT transcendental
with respect to F, since F contains Pi. with respect to F, since F contains Pi.
*/ */
@ -109,12 +109,12 @@ namespace realclosure {
\brief Return the sign of a. \brief Return the sign of a.
*/ */
int sign(numeral const & a); int sign(numeral const & a);
/** /**
\brief Return true if a is zero. \brief Return true if a is zero.
*/ */
bool is_zero(numeral const & a); bool is_zero(numeral const & a);
/** /**
\brief Return true if a is positive. \brief Return true if a is positive.
*/ */
@ -129,13 +129,33 @@ namespace realclosure {
\brief Return true if a is an integer. \brief Return true if a is an integer.
*/ */
bool is_int(numeral const & a); bool is_int(numeral const & a);
/**
\brief Return true if a is a rational.
*/
bool is_rational(numeral const & a);
/**
\brief Return true if a is an algebraic number.
*/
bool is_algebraic(numeral const & a);
/**
\brief Return true if a represents an infinitesimal.
*/
bool is_infinitesimal(numeral const & a);
/**
\brief Return true if a is a transcendental.
*/
bool is_transcendental(numeral const & a);
/** /**
\brief Return true if the representation of \c a depends on \brief Return true if the representation of \c a depends on
infinitesimal extensions. infinitesimal extensions.
*/ */
bool depends_on_infinitesimals(numeral const & a); bool depends_on_infinitesimals(numeral const & a);
/** /**
\brief a <- n \brief a <- n
*/ */
@ -148,14 +168,14 @@ namespace realclosure {
/** /**
\brief Return a^{1/k} \brief Return a^{1/k}
Throws an exception if (a is negative and k is even) or (k is zero). Throws an exception if (a is negative and k is even) or (k is zero).
*/ */
void root(numeral const & a, unsigned k, numeral & b); void root(numeral const & a, unsigned k, numeral & b);
/** /**
\brief Return a^k \brief Return a^k
Throws an exception if 0^0. Throws an exception if 0^0.
*/ */
void power(numeral const & a, unsigned k, numeral & b); void power(numeral const & a, unsigned k, numeral & b);
@ -180,7 +200,7 @@ namespace realclosure {
\brief a <- -a \brief a <- -a
*/ */
void neg(numeral & a); void neg(numeral & a);
/** /**
\brief b <- -a \brief b <- -a
*/ */
@ -190,7 +210,7 @@ namespace realclosure {
\brief a <- 1/a if a != 0 \brief a <- 1/a if a != 0
*/ */
void inv(numeral & a); void inv(numeral & a);
/** /**
\brief b <- 1/a if a != 0 \brief b <- 1/a if a != 0
*/ */
@ -207,7 +227,7 @@ namespace realclosure {
Return 1 if a > b Return 1 if a > b
*/ */
int compare(numeral const & a, numeral const & b); int compare(numeral const & a, numeral const & b);
/** /**
\brief a == b \brief a == b
*/ */
@ -249,7 +269,7 @@ namespace realclosure {
bool ge(numeral const & a, numeral const & b) { return !lt(a, b); } bool ge(numeral const & a, numeral const & b) { return !lt(a, b); }
bool ge(numeral const & a, mpq const & b) { return !lt(a, b); } bool ge(numeral const & a, mpq const & b) { return !lt(a, b); }
bool ge(numeral const & a, mpz const & b) { return !lt(a, b); } bool ge(numeral const & a, mpz const & b) { return !lt(a, b); }
void display(std::ostream & out, numeral const & a, bool compact=false, bool pp=false) const; void display(std::ostream & out, numeral const & a, bool compact=false, bool pp=false) const;
/** /**
@ -259,10 +279,30 @@ namespace realclosure {
*/ */
void display_decimal(std::ostream & out, numeral const & a, unsigned precision = 10) const; void display_decimal(std::ostream & out, numeral const & a, unsigned precision = 10) const;
void display_interval(std::ostream & out, numeral const & a) const; void display_interval(std::ostream & out, numeral const & a) const;
void clean_denominators(numeral const & a, numeral & p, numeral & q); void clean_denominators(numeral const & a, numeral & p, numeral & q);
unsigned extension_index(numeral const & a);
symbol transcendental_name(numeral const &a);
symbol infinitesimal_name(numeral const &a);
unsigned num_coefficients(numeral const &a);
numeral get_coefficient(numeral const &a, unsigned i);
unsigned num_sign_conditions(numeral const &a);
int get_sign_condition_sign(numeral const &a, unsigned i);
bool get_interval(numeral const & a, int & lower_is_inf, int & lower_is_open, numeral & lower, int & upper_is_inf, int & upper_is_open, numeral & upper);
unsigned num_sign_condition_coefficients(numeral const &a, unsigned i);
numeral get_sign_condition_coefficient(numeral const &a, unsigned i, unsigned j);
}; };
struct value; struct value;