mirror of
https://github.com/Z3Prover/z3
synced 2025-06-05 21:53:23 +00:00
Merge branch 'unstable' of https://git01.codeplex.com/z3 into unstable
This commit is contained in:
commit
5d2d89a85c
13 changed files with 333 additions and 200 deletions
|
@ -883,6 +883,29 @@ void incremental_example3() {
|
||||||
std::cout << s.check(a2) << "\n";
|
std::cout << s.check(a2) << "\n";
|
||||||
}
|
}
|
||||||
|
|
||||||
|
void enum_sort_example() {
|
||||||
|
std::cout << "enumeration sort example\n";
|
||||||
|
context ctx;
|
||||||
|
const char * enum_names[] = { "a", "b", "c" };
|
||||||
|
func_decl_vector enum_consts(ctx);
|
||||||
|
func_decl_vector enum_testers(ctx);
|
||||||
|
sort s = ctx.enumeration_sort("enumT", 3, enum_names, enum_consts, enum_testers);
|
||||||
|
// enum_consts[0] is a func_decl of arity 0.
|
||||||
|
// we convert it to an expression using the operator()
|
||||||
|
expr a = enum_consts[0]();
|
||||||
|
expr b = enum_consts[1]();
|
||||||
|
expr x = ctx.constant("x", s);
|
||||||
|
expr test = (x==a) && (x==b);
|
||||||
|
std::cout << "1: " << test << std::endl;
|
||||||
|
tactic qe(ctx, "ctx-solver-simplify");
|
||||||
|
goal g(ctx);
|
||||||
|
g.add(test);
|
||||||
|
expr res(ctx);
|
||||||
|
apply_result result_of_elimination = qe.apply(g);
|
||||||
|
goal result_goal = result_of_elimination[0];
|
||||||
|
std::cout << "2: " << result_goal.as_expr() << std::endl;
|
||||||
|
}
|
||||||
|
|
||||||
int main() {
|
int main() {
|
||||||
try {
|
try {
|
||||||
demorgan(); std::cout << "\n";
|
demorgan(); std::cout << "\n";
|
||||||
|
@ -917,6 +940,7 @@ int main() {
|
||||||
incremental_example1(); std::cout << "\n";
|
incremental_example1(); std::cout << "\n";
|
||||||
incremental_example2(); std::cout << "\n";
|
incremental_example2(); std::cout << "\n";
|
||||||
incremental_example3(); std::cout << "\n";
|
incremental_example3(); std::cout << "\n";
|
||||||
|
enum_sort_example(); std::cout << "\n";
|
||||||
std::cout << "done\n";
|
std::cout << "done\n";
|
||||||
}
|
}
|
||||||
catch (exception & ex) {
|
catch (exception & ex) {
|
||||||
|
|
|
@ -197,7 +197,7 @@ class JavaExample
|
||||||
TestFailedException
|
TestFailedException
|
||||||
{
|
{
|
||||||
Solver s = ctx.mkSolver();
|
Solver s = ctx.mkSolver();
|
||||||
s.assert_(f);
|
s.add(f);
|
||||||
if (s.check() != sat)
|
if (s.check() != sat)
|
||||||
throw new TestFailedException();
|
throw new TestFailedException();
|
||||||
if (sat == Status.SATISFIABLE)
|
if (sat == Status.SATISFIABLE)
|
||||||
|
@ -213,7 +213,7 @@ class JavaExample
|
||||||
System.out.println("\nTactical solver: " + s);
|
System.out.println("\nTactical solver: " + s);
|
||||||
|
|
||||||
for (BoolExpr a : g.getFormulas())
|
for (BoolExpr a : g.getFormulas())
|
||||||
s.assert_(a);
|
s.add(a);
|
||||||
System.out.println("Solver: " + s);
|
System.out.println("Solver: " + s);
|
||||||
|
|
||||||
if (s.check() != sat)
|
if (s.check() != sat)
|
||||||
|
@ -266,8 +266,8 @@ class JavaExample
|
||||||
p.add("mbqi", useMBQI);
|
p.add("mbqi", useMBQI);
|
||||||
s.setParameters(p);
|
s.setParameters(p);
|
||||||
for (BoolExpr a : assumptions)
|
for (BoolExpr a : assumptions)
|
||||||
s.assert_(a);
|
s.add(a);
|
||||||
s.assert_(ctx.mkNot(f));
|
s.add(ctx.mkNot(f));
|
||||||
Status q = s.check();
|
Status q = s.check();
|
||||||
|
|
||||||
switch (q)
|
switch (q)
|
||||||
|
@ -299,8 +299,8 @@ class JavaExample
|
||||||
p.add("mbqi", useMBQI);
|
p.add("mbqi", useMBQI);
|
||||||
s.setParameters(p);
|
s.setParameters(p);
|
||||||
for (BoolExpr a : assumptions)
|
for (BoolExpr a : assumptions)
|
||||||
s.assert_(a);
|
s.add(a);
|
||||||
s.assert_(ctx.mkNot(f));
|
s.add(ctx.mkNot(f));
|
||||||
Status q = s.check();
|
Status q = s.check();
|
||||||
|
|
||||||
switch (q)
|
switch (q)
|
||||||
|
@ -326,9 +326,9 @@ class JavaExample
|
||||||
ArithExpr yr = (ArithExpr) ctx.mkConst(ctx.mkSymbol("y"),
|
ArithExpr yr = (ArithExpr) ctx.mkConst(ctx.mkSymbol("y"),
|
||||||
ctx.mkRealSort());
|
ctx.mkRealSort());
|
||||||
Goal g4 = ctx.mkGoal(true, false, false);
|
Goal g4 = ctx.mkGoal(true, false, false);
|
||||||
g4.assert_(ctx.mkGt(xr, ctx.mkReal(10, 1)));
|
g4.add(ctx.mkGt(xr, ctx.mkReal(10, 1)));
|
||||||
g4.assert_(ctx.mkEq(yr, ctx.mkAdd(xr, ctx.mkReal(1, 1))));
|
g4.add(ctx.mkEq(yr, ctx.mkAdd(xr, ctx.mkReal(1, 1))));
|
||||||
g4.assert_(ctx.mkGt(yr, ctx.mkReal(1, 1)));
|
g4.add(ctx.mkGt(yr, ctx.mkReal(1, 1)));
|
||||||
|
|
||||||
ApplyResult ar = applyTactic(ctx, ctx.mkTactic("simplify"), g4);
|
ApplyResult ar = applyTactic(ctx, ctx.mkTactic("simplify"), g4);
|
||||||
if (ar.getNumSubgoals() == 1
|
if (ar.getNumSubgoals() == 1
|
||||||
|
@ -345,7 +345,7 @@ class JavaExample
|
||||||
|
|
||||||
Solver s = ctx.mkSolver();
|
Solver s = ctx.mkSolver();
|
||||||
for (BoolExpr e : ar.getSubgoals()[0].getFormulas())
|
for (BoolExpr e : ar.getSubgoals()[0].getFormulas())
|
||||||
s.assert_(e);
|
s.add(e);
|
||||||
Status q = s.check();
|
Status q = s.check();
|
||||||
System.out.println("Solver says: " + q);
|
System.out.println("Solver says: " + q);
|
||||||
System.out.println("Model: \n" + s.getModel());
|
System.out.println("Model: \n" + s.getModel());
|
||||||
|
@ -367,7 +367,7 @@ class JavaExample
|
||||||
ctx.mkBitVecSort(32));
|
ctx.mkBitVecSort(32));
|
||||||
ArrayExpr aex = (ArrayExpr) ctx.mkConst(ctx.mkSymbol("MyArray"), asort);
|
ArrayExpr aex = (ArrayExpr) ctx.mkConst(ctx.mkSymbol("MyArray"), asort);
|
||||||
Expr sel = ctx.mkSelect(aex, ctx.mkInt(0));
|
Expr sel = ctx.mkSelect(aex, ctx.mkInt(0));
|
||||||
g.assert_(ctx.mkEq(sel, ctx.mkBV(42, 32)));
|
g.add(ctx.mkEq(sel, ctx.mkBV(42, 32)));
|
||||||
Symbol xs = ctx.mkSymbol("x");
|
Symbol xs = ctx.mkSymbol("x");
|
||||||
IntExpr xc = (IntExpr) ctx.mkConst(xs, ctx.getIntSort());
|
IntExpr xc = (IntExpr) ctx.mkConst(xs, ctx.getIntSort());
|
||||||
|
|
||||||
|
@ -377,11 +377,11 @@ class JavaExample
|
||||||
Expr[] fargs = { ctx.mkConst(xs, ctx.getIntSort()) };
|
Expr[] fargs = { ctx.mkConst(xs, ctx.getIntSort()) };
|
||||||
IntExpr fapp = (IntExpr) ctx.mkApp(fd, fargs);
|
IntExpr fapp = (IntExpr) ctx.mkApp(fd, fargs);
|
||||||
|
|
||||||
g.assert_(ctx.mkEq(ctx.mkAdd(xc, fapp), ctx.mkInt(123)));
|
g.add(ctx.mkEq(ctx.mkAdd(xc, fapp), ctx.mkInt(123)));
|
||||||
|
|
||||||
Solver s = ctx.mkSolver();
|
Solver s = ctx.mkSolver();
|
||||||
for (BoolExpr a : g.getFormulas())
|
for (BoolExpr a : g.getFormulas())
|
||||||
s.assert_(a);
|
s.add(a);
|
||||||
System.out.println("Solver: " + s);
|
System.out.println("Solver: " + s);
|
||||||
|
|
||||||
Status q = s.check();
|
Status q = s.check();
|
||||||
|
@ -574,8 +574,8 @@ class JavaExample
|
||||||
ctx.mkEq(X[i][j], ctx.mkInt(instance[i][j]))));
|
ctx.mkEq(X[i][j], ctx.mkInt(instance[i][j]))));
|
||||||
|
|
||||||
Solver s = ctx.mkSolver();
|
Solver s = ctx.mkSolver();
|
||||||
s.assert_(sudoku_c);
|
s.add(sudoku_c);
|
||||||
s.assert_(instance_c);
|
s.add(instance_c);
|
||||||
|
|
||||||
if (s.check() == Status.SATISFIABLE)
|
if (s.check() == Status.SATISFIABLE)
|
||||||
{
|
{
|
||||||
|
@ -798,14 +798,14 @@ class JavaExample
|
||||||
BoolExpr nontrivial_eq = ctx.mkEq(fapp, fapp2);
|
BoolExpr nontrivial_eq = ctx.mkEq(fapp, fapp2);
|
||||||
|
|
||||||
Goal g = ctx.mkGoal(true, false, false);
|
Goal g = ctx.mkGoal(true, false, false);
|
||||||
g.assert_(trivial_eq);
|
g.add(trivial_eq);
|
||||||
g.assert_(nontrivial_eq);
|
g.add(nontrivial_eq);
|
||||||
System.out.println("Goal: " + g);
|
System.out.println("Goal: " + g);
|
||||||
|
|
||||||
Solver solver = ctx.mkSolver();
|
Solver solver = ctx.mkSolver();
|
||||||
|
|
||||||
for (BoolExpr a : g.getFormulas())
|
for (BoolExpr a : g.getFormulas())
|
||||||
solver.assert_(a);
|
solver.add(a);
|
||||||
|
|
||||||
if (solver.check() != Status.SATISFIABLE)
|
if (solver.check() != Status.SATISFIABLE)
|
||||||
throw new TestFailedException();
|
throw new TestFailedException();
|
||||||
|
@ -820,7 +820,7 @@ class JavaExample
|
||||||
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedSat())
|
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedSat())
|
||||||
throw new TestFailedException();
|
throw new TestFailedException();
|
||||||
|
|
||||||
g.assert_(ctx.mkEq(ctx.mkNumeral(1, ctx.mkBitVecSort(32)),
|
g.add(ctx.mkEq(ctx.mkNumeral(1, ctx.mkBitVecSort(32)),
|
||||||
ctx.mkNumeral(2, ctx.mkBitVecSort(32))));
|
ctx.mkNumeral(2, ctx.mkBitVecSort(32))));
|
||||||
ar = applyTactic(ctx, ctx.mkTactic("smt"), g);
|
ar = applyTactic(ctx, ctx.mkTactic("smt"), g);
|
||||||
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedUnsat())
|
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedUnsat())
|
||||||
|
@ -832,7 +832,7 @@ class JavaExample
|
||||||
throw new TestFailedException();
|
throw new TestFailedException();
|
||||||
|
|
||||||
g2 = ctx.mkGoal(true, true, false);
|
g2 = ctx.mkGoal(true, true, false);
|
||||||
g2.assert_(ctx.mkFalse());
|
g2.add(ctx.mkFalse());
|
||||||
ar = applyTactic(ctx, ctx.mkTactic("smt"), g2);
|
ar = applyTactic(ctx, ctx.mkTactic("smt"), g2);
|
||||||
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedUnsat())
|
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedUnsat())
|
||||||
throw new TestFailedException();
|
throw new TestFailedException();
|
||||||
|
@ -840,10 +840,10 @@ class JavaExample
|
||||||
Goal g3 = ctx.mkGoal(true, true, false);
|
Goal g3 = ctx.mkGoal(true, true, false);
|
||||||
Expr xc = ctx.mkConst(ctx.mkSymbol("x"), ctx.getIntSort());
|
Expr xc = ctx.mkConst(ctx.mkSymbol("x"), ctx.getIntSort());
|
||||||
Expr yc = ctx.mkConst(ctx.mkSymbol("y"), ctx.getIntSort());
|
Expr yc = ctx.mkConst(ctx.mkSymbol("y"), ctx.getIntSort());
|
||||||
g3.assert_(ctx.mkEq(xc, ctx.mkNumeral(1, ctx.getIntSort())));
|
g3.add(ctx.mkEq(xc, ctx.mkNumeral(1, ctx.getIntSort())));
|
||||||
g3.assert_(ctx.mkEq(yc, ctx.mkNumeral(2, ctx.getIntSort())));
|
g3.add(ctx.mkEq(yc, ctx.mkNumeral(2, ctx.getIntSort())));
|
||||||
BoolExpr constr = ctx.mkEq(xc, yc);
|
BoolExpr constr = ctx.mkEq(xc, yc);
|
||||||
g3.assert_(constr);
|
g3.add(constr);
|
||||||
ar = applyTactic(ctx, ctx.mkTactic("smt"), g3);
|
ar = applyTactic(ctx, ctx.mkTactic("smt"), g3);
|
||||||
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedUnsat())
|
if (ar.getNumSubgoals() != 1 || !ar.getSubgoals()[0].isDecidedUnsat())
|
||||||
throw new TestFailedException();
|
throw new TestFailedException();
|
||||||
|
@ -1110,13 +1110,13 @@ class JavaExample
|
||||||
|
|
||||||
// Use a solver for QF_BV
|
// Use a solver for QF_BV
|
||||||
Solver s = ctx.mkSolver("QF_BV");
|
Solver s = ctx.mkSolver("QF_BV");
|
||||||
s.assert_(eq);
|
s.add(eq);
|
||||||
Status res = s.check();
|
Status res = s.check();
|
||||||
System.out.println("solver result: " + res);
|
System.out.println("solver result: " + res);
|
||||||
|
|
||||||
// Or perhaps a tactic for QF_BV
|
// Or perhaps a tactic for QF_BV
|
||||||
Goal g = ctx.mkGoal(true, false, false);
|
Goal g = ctx.mkGoal(true, false, false);
|
||||||
g.assert_(eq);
|
g.add(eq);
|
||||||
|
|
||||||
Tactic t = ctx.mkTactic("qfbv");
|
Tactic t = ctx.mkTactic("qfbv");
|
||||||
ApplyResult ar = t.apply(g);
|
ApplyResult ar = t.apply(g);
|
||||||
|
@ -1139,7 +1139,7 @@ class JavaExample
|
||||||
BoolExpr q = ctx.mkEq(x, y);
|
BoolExpr q = ctx.mkEq(x, y);
|
||||||
|
|
||||||
Goal g = ctx.mkGoal(true, false, false);
|
Goal g = ctx.mkGoal(true, false, false);
|
||||||
g.assert_(q);
|
g.add(q);
|
||||||
|
|
||||||
Tactic t1 = ctx.mkTactic("qfbv");
|
Tactic t1 = ctx.mkTactic("qfbv");
|
||||||
Tactic t2 = ctx.mkTactic("qfbv");
|
Tactic t2 = ctx.mkTactic("qfbv");
|
||||||
|
@ -1341,7 +1341,7 @@ class JavaExample
|
||||||
/* assert x >= "big number" */
|
/* assert x >= "big number" */
|
||||||
BoolExpr c1 = ctx.mkGe(x, big_number);
|
BoolExpr c1 = ctx.mkGe(x, big_number);
|
||||||
System.out.println("assert: x >= 'big number'");
|
System.out.println("assert: x >= 'big number'");
|
||||||
solver.assert_(c1);
|
solver.add(c1);
|
||||||
|
|
||||||
/* create a backtracking point */
|
/* create a backtracking point */
|
||||||
System.out.println("push");
|
System.out.println("push");
|
||||||
|
@ -1350,7 +1350,7 @@ class JavaExample
|
||||||
/* assert x <= 3 */
|
/* assert x <= 3 */
|
||||||
BoolExpr c2 = ctx.mkLe(x, three);
|
BoolExpr c2 = ctx.mkLe(x, three);
|
||||||
System.out.println("assert: x <= 3");
|
System.out.println("assert: x <= 3");
|
||||||
solver.assert_(c2);
|
solver.add(c2);
|
||||||
|
|
||||||
/* context is inconsistent at this point */
|
/* context is inconsistent at this point */
|
||||||
if (solver.check() != Status.UNSATISFIABLE)
|
if (solver.check() != Status.UNSATISFIABLE)
|
||||||
|
@ -1375,7 +1375,7 @@ class JavaExample
|
||||||
/* assert y > x */
|
/* assert y > x */
|
||||||
BoolExpr c3 = ctx.mkGt(y, x);
|
BoolExpr c3 = ctx.mkGt(y, x);
|
||||||
System.out.println("assert: y > x");
|
System.out.println("assert: y > x");
|
||||||
solver.assert_(c3);
|
solver.add(c3);
|
||||||
|
|
||||||
/* the context is still consistent. */
|
/* the context is still consistent. */
|
||||||
if (solver.check() != Status.SATISFIABLE)
|
if (solver.check() != Status.SATISFIABLE)
|
||||||
|
@ -1911,10 +1911,10 @@ class JavaExample
|
||||||
Solver solver = ctx.mkSolver();
|
Solver solver = ctx.mkSolver();
|
||||||
|
|
||||||
/* assert x < y */
|
/* assert x < y */
|
||||||
solver.assert_(ctx.mkLt(x, y));
|
solver.add(ctx.mkLt(x, y));
|
||||||
|
|
||||||
/* assert x > 2 */
|
/* assert x > 2 */
|
||||||
solver.assert_(ctx.mkGt(x, two));
|
solver.add(ctx.mkGt(x, two));
|
||||||
|
|
||||||
/* find model for the constraints above */
|
/* find model for the constraints above */
|
||||||
Model model = null;
|
Model model = null;
|
||||||
|
@ -1964,9 +1964,9 @@ class JavaExample
|
||||||
Solver solver = ctx.mkSolver();
|
Solver solver = ctx.mkSolver();
|
||||||
|
|
||||||
/* assert tup1 != tup2 */
|
/* assert tup1 != tup2 */
|
||||||
solver.assert_(ctx.mkNot(ctx.mkEq(tup1, tup2)));
|
solver.add(ctx.mkNot(ctx.mkEq(tup1, tup2)));
|
||||||
/* assert first tup1 = first tup2 */
|
/* assert first tup1 = first tup2 */
|
||||||
solver.assert_(ctx.mkEq(ctx.mkApp(first, tup1), ctx.mkApp(first, tup2)));
|
solver.add(ctx.mkEq(ctx.mkApp(first, tup1), ctx.mkApp(first, tup2)));
|
||||||
|
|
||||||
/* find model for the constraints above */
|
/* find model for the constraints above */
|
||||||
Model model = null;
|
Model model = null;
|
||||||
|
@ -2014,7 +2014,7 @@ class JavaExample
|
||||||
// Assert all feasible bounds.
|
// Assert all feasible bounds.
|
||||||
for (int i = 0; i < num_Exprs; ++i)
|
for (int i = 0; i < num_Exprs; ++i)
|
||||||
{
|
{
|
||||||
solver.assert_(ctx.mkBVULE(to_minimize[i],
|
solver.add(ctx.mkBVULE(to_minimize[i],
|
||||||
ctx.mkBV(upper[i], 32)));
|
ctx.mkBV(upper[i], 32)));
|
||||||
}
|
}
|
||||||
|
|
||||||
|
@ -2050,7 +2050,7 @@ class JavaExample
|
||||||
{
|
{
|
||||||
last_upper = (upper[i] + lower[i]) / 2;
|
last_upper = (upper[i] + lower[i]) / 2;
|
||||||
last_index = i;
|
last_index = i;
|
||||||
solver.assert_(ctx.mkBVULE(to_minimize[i],
|
solver.add(ctx.mkBVULE(to_minimize[i],
|
||||||
ctx.mkBV(last_upper, 32)));
|
ctx.mkBV(last_upper, 32)));
|
||||||
some_work = true;
|
some_work = true;
|
||||||
break;
|
break;
|
||||||
|
@ -2074,7 +2074,7 @@ class JavaExample
|
||||||
|
|
||||||
Solver solver = ctx.mkSolver();
|
Solver solver = ctx.mkSolver();
|
||||||
|
|
||||||
solver.assert_(ctx.mkBVULE(x, ctx.mkBVAdd(y, z)));
|
solver.add(ctx.mkBVULE(x, ctx.mkBVAdd(y, z)));
|
||||||
checkSmall(ctx, solver, x, y, z);
|
checkSmall(ctx, solver, x, y, z);
|
||||||
}
|
}
|
||||||
|
|
||||||
|
@ -2120,10 +2120,10 @@ class JavaExample
|
||||||
BoolExpr f3 = ctx.mkOr(ctx.mkNot(pa), ctx.mkNot(pc));
|
BoolExpr f3 = ctx.mkOr(ctx.mkNot(pa), ctx.mkNot(pc));
|
||||||
BoolExpr f4 = pd;
|
BoolExpr f4 = pd;
|
||||||
|
|
||||||
solver.assert_(ctx.mkOr(f1, p1));
|
solver.add(ctx.mkOr(f1, p1));
|
||||||
solver.assert_(ctx.mkOr(f2, p2));
|
solver.add(ctx.mkOr(f2, p2));
|
||||||
solver.assert_(ctx.mkOr(f3, p3));
|
solver.add(ctx.mkOr(f3, p3));
|
||||||
solver.assert_(ctx.mkOr(f4, p4));
|
solver.add(ctx.mkOr(f4, p4));
|
||||||
Status result = solver.check(assumptions);
|
Status result = solver.check(assumptions);
|
||||||
|
|
||||||
if (result == Status.UNSATISFIABLE)
|
if (result == Status.UNSATISFIABLE)
|
||||||
|
|
|
@ -30,13 +30,12 @@ CC=getenv("CC", None)
|
||||||
CPPFLAGS=getenv("CPPFLAGS", "")
|
CPPFLAGS=getenv("CPPFLAGS", "")
|
||||||
CXXFLAGS=getenv("CXXFLAGS", "")
|
CXXFLAGS=getenv("CXXFLAGS", "")
|
||||||
LDFLAGS=getenv("LDFLAGS", "")
|
LDFLAGS=getenv("LDFLAGS", "")
|
||||||
JAVA=getenv("JAVA", "java")
|
|
||||||
JAVAC=getenv("JAVAC", "javac")
|
|
||||||
JAVA_HOME=getenv("JAVA_HOME", None)
|
|
||||||
JNI_HOME=getenv("JNI_HOME", None)
|
JNI_HOME=getenv("JNI_HOME", None)
|
||||||
|
|
||||||
CXX_COMPILERS=['g++', 'clang++']
|
CXX_COMPILERS=['g++', 'clang++']
|
||||||
C_COMPILERS=['gcc', 'clang']
|
C_COMPILERS=['gcc', 'clang']
|
||||||
|
JAVAC=None
|
||||||
|
JAR=None
|
||||||
PYTHON_PACKAGE_DIR=distutils.sysconfig.get_python_lib()
|
PYTHON_PACKAGE_DIR=distutils.sysconfig.get_python_lib()
|
||||||
BUILD_DIR='build'
|
BUILD_DIR='build'
|
||||||
REV_BUILD_DIR='..'
|
REV_BUILD_DIR='..'
|
||||||
|
@ -207,24 +206,6 @@ def test_openmp(cc):
|
||||||
t.commit()
|
t.commit()
|
||||||
return exec_compiler_cmd([cc, CPPFLAGS, 'tstomp.cpp', LDFLAGS, '-fopenmp']) == 0
|
return exec_compiler_cmd([cc, CPPFLAGS, 'tstomp.cpp', LDFLAGS, '-fopenmp']) == 0
|
||||||
|
|
||||||
def check_java():
|
|
||||||
t = TempFile('Hello.java')
|
|
||||||
t.add('public class Hello { public static void main(String[] args) { System.out.println("Hello, World"); }}\n')
|
|
||||||
t.commit()
|
|
||||||
if is_verbose():
|
|
||||||
print("Testing %s..." % JAVAC)
|
|
||||||
r = exec_cmd([JAVAC, 'Hello.java'])
|
|
||||||
if r != 0:
|
|
||||||
raise MKException('Failed testing Java compiler. Set environment variable JAVAC with the path to the Java compiler')
|
|
||||||
if is_verbose():
|
|
||||||
print("Testing %s..." % JAVA)
|
|
||||||
r = exec_cmd([JAVA, 'Hello'])
|
|
||||||
rmf('Hello.class')
|
|
||||||
if r != 0:
|
|
||||||
raise MKException('Failed testing Java program. Set environment variable JAVA with the path to the Java virtual machine')
|
|
||||||
find_java_home()
|
|
||||||
find_jni_home()
|
|
||||||
|
|
||||||
def find_jni_h(path):
|
def find_jni_h(path):
|
||||||
for root, dirs, files in os.walk(path):
|
for root, dirs, files in os.walk(path):
|
||||||
for f in files:
|
for f in files:
|
||||||
|
@ -232,76 +213,110 @@ def find_jni_h(path):
|
||||||
return root
|
return root
|
||||||
return False
|
return False
|
||||||
|
|
||||||
def find_java_home():
|
def check_java():
|
||||||
global JAVA_HOME
|
|
||||||
if JAVA_HOME != None:
|
|
||||||
if IS_WINDOWS:
|
|
||||||
ind = '%s%s' % (JAVA_HOME, '\\bin\\java.exe')
|
|
||||||
else:
|
|
||||||
ind = '%s%s' % (JAVA_HOME, '/bin/java')
|
|
||||||
if not os.path.exists(ind):
|
|
||||||
raise MKException("Failed to detect java at '%s'.Possible solution: set JAVA_HOME with the path to JDK." % os.path.join(JAVA_HOME))
|
|
||||||
else:
|
|
||||||
return
|
|
||||||
if is_verbose():
|
|
||||||
print("Finding JAVA_HOME...")
|
|
||||||
t = TempFile('output')
|
|
||||||
null = open(os.devnull, 'wb')
|
|
||||||
try:
|
|
||||||
subprocess.call([JAVA, '-verbose'], stdout=t.fname, stderr=null)
|
|
||||||
t.commit()
|
|
||||||
except:
|
|
||||||
raise MKException('Failed to find JAVA_HOME')
|
|
||||||
open_pat = re.compile("\[Opened (.*)\]")
|
|
||||||
t = open('output', 'r')
|
|
||||||
for line in t:
|
|
||||||
m = open_pat.match(line)
|
|
||||||
if m:
|
|
||||||
# Remove last 3 directives from m.group(1)
|
|
||||||
tmp = m.group(1).split(os.sep)
|
|
||||||
path = string.join(tmp[:len(tmp) - 3], os.sep)
|
|
||||||
if IS_WINDOWS:
|
|
||||||
ind = '%s%s' % (path, '\\bin\\java.exe')
|
|
||||||
else:
|
|
||||||
ind = '%s%s' % (path, '/bin/java')
|
|
||||||
if os.path.exists(ind):
|
|
||||||
JAVA_HOME = path
|
|
||||||
return
|
|
||||||
if IS_OSX:
|
|
||||||
path = '%s%s' % (path, '/Contents/Home/')
|
|
||||||
ind = '%s%s' % (path, 'bin/java')
|
|
||||||
if os.path.exists(ind):
|
|
||||||
JAVA_HOME = path
|
|
||||||
return
|
|
||||||
raise MKException("Failed to detect java at '%s'.Possible solution: set JAVA_HOME with the path to JDK." % os.path.join(path))
|
|
||||||
return
|
|
||||||
raise MKException('Failed to find JAVA_HOME')
|
|
||||||
|
|
||||||
def find_jni_home():
|
|
||||||
global JNI_HOME
|
global JNI_HOME
|
||||||
if JNI_HOME != None:
|
global JAVAC
|
||||||
if is_verbose():
|
global JAR
|
||||||
print("Checking jni.h...")
|
|
||||||
path = JNI_HOME
|
JDK_HOME = getenv('JDK_HOME', None) # we only need to check this locally.
|
||||||
fn = os.path.join(path, 'jni.h')
|
|
||||||
print("Checking for jni.h in %s..." % JNI_HOME)
|
if is_verbose():
|
||||||
if os.path.exists(fn):
|
print("Finding javac ...")
|
||||||
return
|
|
||||||
|
if JDK_HOME != None:
|
||||||
|
if IS_WINDOWS:
|
||||||
|
JAVAC = os.path.join(JDK_HOME, 'bin', 'javac.exe')
|
||||||
|
else:
|
||||||
|
JAVAC = os.path.join(JDK_HOME, 'bin', 'javac')
|
||||||
|
|
||||||
|
if not os.path.exists(JAVAC):
|
||||||
|
raise MKException("Failed to detect javac at '%s/bin'; the environment variable JDK_HOME is probably set to the wrong path." % os.path.join(JDK_HOME))
|
||||||
else:
|
else:
|
||||||
path = '%s%s' % (JAVA_HOME, '/include/')
|
# Search for javac in the path.
|
||||||
fn = '%s%s' % (path, 'jni.h')
|
ind = 'javac';
|
||||||
print("Checking for jni.h in %s..." % path)
|
if IS_WINDOWS:
|
||||||
if os.path.exists(fn):
|
ind = ind + '.exe'
|
||||||
JNI_HOME = find_jni_h(path)
|
paths = os.getenv('PATH', None)
|
||||||
elif IS_OSX:
|
if paths:
|
||||||
# Apparently Apple knows best where to put stuff...
|
spaths = paths.split(os.pathsep)
|
||||||
path = '/System/Library/Frameworks/JavaVM.framework/Headers/'
|
for i in range(0, len(spaths)):
|
||||||
fn = '%s%s' % (path, 'jni.h')
|
cmb = os.path.join(spaths[i], ind)
|
||||||
print("Checking for jni.h in %s..." % path)
|
if os.path.exists(cmb):
|
||||||
if os.path.exists(fn):
|
JAVAC = cmb
|
||||||
JNI_HOME = find_jni_h(path)
|
break
|
||||||
if JNI_HOME == None:
|
|
||||||
raise MKException("Failed to detect jni.h. Possible solution: set JNI_HOME with the path to JDK.")
|
if JAVAC == None:
|
||||||
|
raise MKException('No java compiler in the path, please adjust your PATH or set JDK_HOME to the location of the JDK.')
|
||||||
|
|
||||||
|
if is_verbose():
|
||||||
|
print("Finding jar ...")
|
||||||
|
|
||||||
|
if IS_WINDOWS:
|
||||||
|
JAR = os.path.join(os.path.dirname(JAVAC), 'jar.exe')
|
||||||
|
else:
|
||||||
|
JAR = os.path.join(os.path.dirname(JAVAC), 'jar')
|
||||||
|
|
||||||
|
if not os.path.exists(JAR):
|
||||||
|
raise MKException("Failed to detect jar at '%s'; the environment variable JDK_HOME is probably set to the wrong path." % os.path.join(JDK_HOME))
|
||||||
|
|
||||||
|
if is_verbose():
|
||||||
|
print("Testing %s..." % JAVAC)
|
||||||
|
|
||||||
|
t = TempFile('Hello.java')
|
||||||
|
t.add('public class Hello { public static void main(String[] args) { System.out.println("Hello, World"); }}\n')
|
||||||
|
t.commit()
|
||||||
|
|
||||||
|
oo = TempFile('output')
|
||||||
|
eo = TempFile('errout')
|
||||||
|
try:
|
||||||
|
subprocess.call([JAVAC, 'Hello.java', '-verbose'], stdout=oo.fname, stderr=eo.fname)
|
||||||
|
oo.commit()
|
||||||
|
eo.commit()
|
||||||
|
except:
|
||||||
|
raise MKException('Found, but failed to run Java compiler at %s' % (JAVAC))
|
||||||
|
|
||||||
|
os.remove('Hello.class')
|
||||||
|
|
||||||
|
if is_verbose():
|
||||||
|
print("Finding jni.h...")
|
||||||
|
|
||||||
|
if JNI_HOME != None:
|
||||||
|
if not os.path.exists(path.join(JNI_HOME, 'jni.h')):
|
||||||
|
raise MKException("Failed to detect jni.h '%s'; the environment variable JNI_HOME is probably set to the wrong path." % os.path.join(JNI_HOME))
|
||||||
|
else:
|
||||||
|
# Search for jni.h in the library directories...
|
||||||
|
t = open('errout', 'r')
|
||||||
|
open_pat = re.compile("\[search path for class files: (.*)\]")
|
||||||
|
cdirs = []
|
||||||
|
for line in t:
|
||||||
|
m = open_pat.match(line)
|
||||||
|
if m:
|
||||||
|
libdirs = m.group(1).split(',')
|
||||||
|
for libdir in libdirs:
|
||||||
|
q = os.path.dirname(libdir)
|
||||||
|
if cdirs.count(q) == 0:
|
||||||
|
cdirs.append(q)
|
||||||
|
|
||||||
|
# ... plus some heuristic ones.
|
||||||
|
extra_dirs = []
|
||||||
|
|
||||||
|
# For the libraries, even the JDK usually uses a JRE that comes with it. To find the
|
||||||
|
# headers we have to go a little bit higher up.
|
||||||
|
for dir in cdirs:
|
||||||
|
extra_dirs.append(os.path.abspath(os.path.join(dir, '..')))
|
||||||
|
|
||||||
|
if IS_OSX: # Apparently Apple knows best where to put stuff...
|
||||||
|
extra_dirs.append('/System/Library/Frameworks/JavaVM.framework/Headers/')
|
||||||
|
|
||||||
|
cdirs[len(cdirs):] = extra_dirs
|
||||||
|
|
||||||
|
for dir in cdirs:
|
||||||
|
q = find_jni_h(dir)
|
||||||
|
if q != False:
|
||||||
|
JNI_HOME = q
|
||||||
|
|
||||||
|
if JNI_HOME == None:
|
||||||
|
raise MKException("Failed to detect jni.h. Possible solution: set JNI_HOME with the path to JDK.")
|
||||||
|
|
||||||
def is64():
|
def is64():
|
||||||
return sys.maxsize >= 2**32
|
return sys.maxsize >= 2**32
|
||||||
|
@ -436,9 +451,7 @@ def display_help(exit_code):
|
||||||
print(" LDFLAGS Linker flags, e.g., -L<lib dir> if you have libraries in a non-standard directory")
|
print(" LDFLAGS Linker flags, e.g., -L<lib dir> if you have libraries in a non-standard directory")
|
||||||
print(" CPPFLAGS Preprocessor flags, e.g., -I<include dir> if you have header files in a non-standard directory")
|
print(" CPPFLAGS Preprocessor flags, e.g., -I<include dir> if you have header files in a non-standard directory")
|
||||||
print(" CXXFLAGS C++ compiler flags")
|
print(" CXXFLAGS C++ compiler flags")
|
||||||
print(" JAVA Java virtual machine (only relevant if -j or --java option is provided)")
|
print(" JDK_HOME JDK installation directory (only relevant if -j or --java option is provided)")
|
||||||
print(" JAVAC Java compiler (only relevant if -j or --java option is provided)")
|
|
||||||
print(" JAVA_HOME JDK installation directory (only relevant if -j or --java option is provided)")
|
|
||||||
print(" JNI_HOME JNI bindings directory (only relevant if -j or --java option is provided)")
|
print(" JNI_HOME JNI bindings directory (only relevant if -j or --java option is provided)")
|
||||||
exit(exit_code)
|
exit(exit_code)
|
||||||
|
|
||||||
|
@ -1120,6 +1133,9 @@ class JavaDLLComponent(Component):
|
||||||
self.manifest_file = manifest_file
|
self.manifest_file = manifest_file
|
||||||
|
|
||||||
def mk_makefile(self, out):
|
def mk_makefile(self, out):
|
||||||
|
global JAVAC
|
||||||
|
global JAR
|
||||||
|
|
||||||
if is_java_enabled():
|
if is_java_enabled():
|
||||||
mk_dir(os.path.join(BUILD_DIR, 'api', 'java', 'classes'))
|
mk_dir(os.path.join(BUILD_DIR, 'api', 'java', 'classes'))
|
||||||
dllfile = '%s$(SO_EXT)' % self.dll_name
|
dllfile = '%s$(SO_EXT)' % self.dll_name
|
||||||
|
@ -1146,6 +1162,9 @@ class JavaDLLComponent(Component):
|
||||||
deps += '%s ' % os.path.join(self.to_src_dir, 'enumerations', jfile)
|
deps += '%s ' % os.path.join(self.to_src_dir, 'enumerations', jfile)
|
||||||
out.write(deps)
|
out.write(deps)
|
||||||
out.write('\n')
|
out.write('\n')
|
||||||
|
if IS_WINDOWS:
|
||||||
|
JAVAC = '"%s"' % JAVAC
|
||||||
|
JAR = '"%s"' % JAR
|
||||||
t = ('\t%s %s.java -d %s\n' % (JAVAC, os.path.join(self.to_src_dir, 'enumerations', '*'), os.path.join('api', 'java', 'classes')))
|
t = ('\t%s %s.java -d %s\n' % (JAVAC, os.path.join(self.to_src_dir, 'enumerations', '*'), os.path.join('api', 'java', 'classes')))
|
||||||
out.write(t)
|
out.write(t)
|
||||||
t = ('\t%s -cp %s %s.java -d %s\n' % (JAVAC,
|
t = ('\t%s -cp %s %s.java -d %s\n' % (JAVAC,
|
||||||
|
@ -1153,7 +1172,7 @@ class JavaDLLComponent(Component):
|
||||||
os.path.join(self.to_src_dir, '*'),
|
os.path.join(self.to_src_dir, '*'),
|
||||||
os.path.join('api', 'java', 'classes')))
|
os.path.join('api', 'java', 'classes')))
|
||||||
out.write(t)
|
out.write(t)
|
||||||
out.write('\tjar cfm %s.jar %s -C %s .\n' % (self.package_name,
|
out.write('\t%s cfm %s.jar %s -C %s .\n' % (JAR, self.package_name,
|
||||||
os.path.join(self.to_src_dir, 'manifest'),
|
os.path.join(self.to_src_dir, 'manifest'),
|
||||||
os.path.join('api', 'java', 'classes')))
|
os.path.join('api', 'java', 'classes')))
|
||||||
out.write('java: %s.jar\n\n' % self.package_name)
|
out.write('java: %s.jar\n\n' % self.package_name)
|
||||||
|
@ -1425,10 +1444,8 @@ def mk_config():
|
||||||
if is_verbose():
|
if is_verbose():
|
||||||
print('64-bit: %s' % is64())
|
print('64-bit: %s' % is64())
|
||||||
if is_java_enabled():
|
if is_java_enabled():
|
||||||
print('Java Home: %s' % JAVA_HOME)
|
print('JNI Bindings: %s' % JNI_HOME)
|
||||||
print('JNI Home: %s' % JNI_HOME)
|
|
||||||
print('Java Compiler: %s' % JAVAC)
|
print('Java Compiler: %s' % JAVAC)
|
||||||
print('Java VM: %s' % JAVA)
|
|
||||||
else:
|
else:
|
||||||
global CXX, CC, GMP, CPPFLAGS, CXXFLAGS, LDFLAGS
|
global CXX, CC, GMP, CPPFLAGS, CXXFLAGS, LDFLAGS
|
||||||
ARITH = "internal"
|
ARITH = "internal"
|
||||||
|
@ -1530,10 +1547,8 @@ def mk_config():
|
||||||
print('gprof: enabled')
|
print('gprof: enabled')
|
||||||
print('Python version: %s' % distutils.sysconfig.get_python_version())
|
print('Python version: %s' % distutils.sysconfig.get_python_version())
|
||||||
if is_java_enabled():
|
if is_java_enabled():
|
||||||
print('Java Home: %s' % JAVA_HOME)
|
print('JNI Bindings: %s' % JNI_HOME)
|
||||||
print('JNI Home: %s' % JNI_HOME)
|
|
||||||
print('Java Compiler: %s' % JAVAC)
|
print('Java Compiler: %s' % JAVAC)
|
||||||
print('Java VM: %s' % JAVA)
|
|
||||||
|
|
||||||
def mk_install(out):
|
def mk_install(out):
|
||||||
out.write('install:\n')
|
out.write('install:\n')
|
||||||
|
|
|
@ -63,6 +63,11 @@ namespace z3 {
|
||||||
class statistics;
|
class statistics;
|
||||||
class apply_result;
|
class apply_result;
|
||||||
class fixedpoint;
|
class fixedpoint;
|
||||||
|
template<typename T> class ast_vector_tpl;
|
||||||
|
typedef ast_vector_tpl<ast> ast_vector;
|
||||||
|
typedef ast_vector_tpl<expr> expr_vector;
|
||||||
|
typedef ast_vector_tpl<sort> sort_vector;
|
||||||
|
typedef ast_vector_tpl<func_decl> func_decl_vector;
|
||||||
|
|
||||||
inline void set_param(char const * param, char const * value) { Z3_global_param_set(param, value); }
|
inline void set_param(char const * param, char const * value) { Z3_global_param_set(param, value); }
|
||||||
inline void set_param(char const * param, bool value) { Z3_global_param_set(param, value ? "true" : "false"); }
|
inline void set_param(char const * param, bool value) { Z3_global_param_set(param, value ? "true" : "false"); }
|
||||||
|
@ -190,7 +195,13 @@ namespace z3 {
|
||||||
Example: Given a context \c c, <tt>c.array_sort(c.int_sort(), c.bool_sort())</tt> is an array sort from integer to Boolean.
|
Example: Given a context \c c, <tt>c.array_sort(c.int_sort(), c.bool_sort())</tt> is an array sort from integer to Boolean.
|
||||||
*/
|
*/
|
||||||
sort array_sort(sort d, sort r);
|
sort array_sort(sort d, sort r);
|
||||||
|
/**
|
||||||
|
\brief Return an enumeration sort: enum_names[0], ..., enum_names[n-1].
|
||||||
|
\c cs and \c ts are output parameters. The method stores in \c cs the constants corresponding to the enumerated elements,
|
||||||
|
and in \c ts the predicates for testing if terms of the enumeration sort correspond to an enumeration.
|
||||||
|
*/
|
||||||
|
sort enumeration_sort(char const * name, unsigned n, char const * const * enum_names, func_decl_vector & cs, func_decl_vector & ts);
|
||||||
|
|
||||||
func_decl function(symbol const & name, unsigned arity, sort const * domain, sort const & range);
|
func_decl function(symbol const & name, unsigned arity, sort const * domain, sort const & range);
|
||||||
func_decl function(char const * name, unsigned arity, sort const * domain, sort const & range);
|
func_decl function(char const * name, unsigned arity, sort const * domain, sort const & range);
|
||||||
func_decl function(char const * name, sort const & domain, sort const & range);
|
func_decl function(char const * name, sort const & domain, sort const & range);
|
||||||
|
@ -240,8 +251,8 @@ namespace z3 {
|
||||||
array(unsigned sz):m_size(sz) { m_array = new T[sz]; }
|
array(unsigned sz):m_size(sz) { m_array = new T[sz]; }
|
||||||
~array() { delete[] m_array; }
|
~array() { delete[] m_array; }
|
||||||
unsigned size() const { return m_size; }
|
unsigned size() const { return m_size; }
|
||||||
T & operator[](unsigned i) { assert(i < m_size); return m_array[i]; }
|
T & operator[](int i) { assert(0 <= i); assert(static_cast<unsigned>(i) < m_size); return m_array[i]; }
|
||||||
T const & operator[](unsigned i) const { assert(i < m_size); return m_array[i]; }
|
T const & operator[](int i) const { assert(0 <= i); assert(static_cast<unsigned>(i) < m_size); return m_array[i]; }
|
||||||
T const * ptr() const { return m_array; }
|
T const * ptr() const { return m_array; }
|
||||||
T * ptr() { return m_array; }
|
T * ptr() { return m_array; }
|
||||||
};
|
};
|
||||||
|
@ -414,6 +425,7 @@ namespace z3 {
|
||||||
|
|
||||||
bool is_const() const { return arity() == 0; }
|
bool is_const() const { return arity() == 0; }
|
||||||
|
|
||||||
|
expr operator()() const;
|
||||||
expr operator()(unsigned n, expr const * args) const;
|
expr operator()(unsigned n, expr const * args) const;
|
||||||
expr operator()(expr const & a) const;
|
expr operator()(expr const & a) const;
|
||||||
expr operator()(int a) const;
|
expr operator()(int a) const;
|
||||||
|
@ -1004,7 +1016,7 @@ namespace z3 {
|
||||||
~ast_vector_tpl() { Z3_ast_vector_dec_ref(ctx(), m_vector); }
|
~ast_vector_tpl() { Z3_ast_vector_dec_ref(ctx(), m_vector); }
|
||||||
operator Z3_ast_vector() const { return m_vector; }
|
operator Z3_ast_vector() const { return m_vector; }
|
||||||
unsigned size() const { return Z3_ast_vector_size(ctx(), m_vector); }
|
unsigned size() const { return Z3_ast_vector_size(ctx(), m_vector); }
|
||||||
T operator[](unsigned i) const { Z3_ast r = Z3_ast_vector_get(ctx(), m_vector, i); check_error(); return cast_ast<T>()(ctx(), r); }
|
T operator[](int i) const { assert(0 <= i); Z3_ast r = Z3_ast_vector_get(ctx(), m_vector, i); check_error(); return cast_ast<T>()(ctx(), r); }
|
||||||
void push_back(T const & e) { Z3_ast_vector_push(ctx(), m_vector, e); check_error(); }
|
void push_back(T const & e) { Z3_ast_vector_push(ctx(), m_vector, e); check_error(); }
|
||||||
void resize(unsigned sz) { Z3_ast_vector_resize(ctx(), m_vector, sz); check_error(); }
|
void resize(unsigned sz) { Z3_ast_vector_resize(ctx(), m_vector, sz); check_error(); }
|
||||||
T back() const { return operator[](size() - 1); }
|
T back() const { return operator[](size() - 1); }
|
||||||
|
@ -1020,11 +1032,6 @@ namespace z3 {
|
||||||
friend std::ostream & operator<<(std::ostream & out, ast_vector_tpl const & v) { out << Z3_ast_vector_to_string(v.ctx(), v); return out; }
|
friend std::ostream & operator<<(std::ostream & out, ast_vector_tpl const & v) { out << Z3_ast_vector_to_string(v.ctx(), v); return out; }
|
||||||
};
|
};
|
||||||
|
|
||||||
typedef ast_vector_tpl<ast> ast_vector;
|
|
||||||
typedef ast_vector_tpl<expr> expr_vector;
|
|
||||||
typedef ast_vector_tpl<sort> sort_vector;
|
|
||||||
typedef ast_vector_tpl<func_decl> func_decl_vector;
|
|
||||||
|
|
||||||
class func_entry : public object {
|
class func_entry : public object {
|
||||||
Z3_func_entry m_entry;
|
Z3_func_entry m_entry;
|
||||||
void init(Z3_func_entry e) {
|
void init(Z3_func_entry e) {
|
||||||
|
@ -1105,7 +1112,10 @@ namespace z3 {
|
||||||
func_decl get_const_decl(unsigned i) const { Z3_func_decl r = Z3_model_get_const_decl(ctx(), m_model, i); check_error(); return func_decl(ctx(), r); }
|
func_decl get_const_decl(unsigned i) const { Z3_func_decl r = Z3_model_get_const_decl(ctx(), m_model, i); check_error(); return func_decl(ctx(), r); }
|
||||||
func_decl get_func_decl(unsigned i) const { Z3_func_decl r = Z3_model_get_func_decl(ctx(), m_model, i); check_error(); return func_decl(ctx(), r); }
|
func_decl get_func_decl(unsigned i) const { Z3_func_decl r = Z3_model_get_func_decl(ctx(), m_model, i); check_error(); return func_decl(ctx(), r); }
|
||||||
unsigned size() const { return num_consts() + num_funcs(); }
|
unsigned size() const { return num_consts() + num_funcs(); }
|
||||||
func_decl operator[](unsigned i) const { return i < num_consts() ? get_const_decl(i) : get_func_decl(i - num_consts()); }
|
func_decl operator[](int i) const {
|
||||||
|
assert(0 <= i);
|
||||||
|
return static_cast<unsigned>(i) < num_consts() ? get_const_decl(i) : get_func_decl(i - num_consts());
|
||||||
|
}
|
||||||
|
|
||||||
expr get_const_interp(func_decl c) const {
|
expr get_const_interp(func_decl c) const {
|
||||||
check_context(*this, c);
|
check_context(*this, c);
|
||||||
|
@ -1253,7 +1263,7 @@ namespace z3 {
|
||||||
}
|
}
|
||||||
void add(expr const & f) { check_context(*this, f); Z3_goal_assert(ctx(), m_goal, f); check_error(); }
|
void add(expr const & f) { check_context(*this, f); Z3_goal_assert(ctx(), m_goal, f); check_error(); }
|
||||||
unsigned size() const { return Z3_goal_size(ctx(), m_goal); }
|
unsigned size() const { return Z3_goal_size(ctx(), m_goal); }
|
||||||
expr operator[](unsigned i) const { Z3_ast r = Z3_goal_formula(ctx(), m_goal, i); check_error(); return expr(ctx(), r); }
|
expr operator[](int i) const { assert(0 <= i); Z3_ast r = Z3_goal_formula(ctx(), m_goal, i); check_error(); return expr(ctx(), r); }
|
||||||
Z3_goal_prec precision() const { return Z3_goal_precision(ctx(), m_goal); }
|
Z3_goal_prec precision() const { return Z3_goal_precision(ctx(), m_goal); }
|
||||||
bool inconsistent() const { return Z3_goal_inconsistent(ctx(), m_goal) != 0; }
|
bool inconsistent() const { return Z3_goal_inconsistent(ctx(), m_goal) != 0; }
|
||||||
unsigned depth() const { return Z3_goal_depth(ctx(), m_goal); }
|
unsigned depth() const { return Z3_goal_depth(ctx(), m_goal); }
|
||||||
|
@ -1261,6 +1271,19 @@ namespace z3 {
|
||||||
unsigned num_exprs() const { return Z3_goal_num_exprs(ctx(), m_goal); }
|
unsigned num_exprs() const { return Z3_goal_num_exprs(ctx(), m_goal); }
|
||||||
bool is_decided_sat() const { return Z3_goal_is_decided_sat(ctx(), m_goal) != 0; }
|
bool is_decided_sat() const { return Z3_goal_is_decided_sat(ctx(), m_goal) != 0; }
|
||||||
bool is_decided_unsat() const { return Z3_goal_is_decided_unsat(ctx(), m_goal) != 0; }
|
bool is_decided_unsat() const { return Z3_goal_is_decided_unsat(ctx(), m_goal) != 0; }
|
||||||
|
expr as_expr() const {
|
||||||
|
unsigned n = size();
|
||||||
|
if (n == 0)
|
||||||
|
return ctx().bool_val(false);
|
||||||
|
else if (n == 1)
|
||||||
|
return operator[](0);
|
||||||
|
else {
|
||||||
|
array<Z3_ast> args(n);
|
||||||
|
for (unsigned i = 0; i < n; i++)
|
||||||
|
args[i] = operator[](i);
|
||||||
|
return expr(ctx(), Z3_mk_and(ctx(), n, args.ptr()));
|
||||||
|
}
|
||||||
|
}
|
||||||
friend std::ostream & operator<<(std::ostream & out, goal const & g) { out << Z3_goal_to_string(g.ctx(), g); return out; }
|
friend std::ostream & operator<<(std::ostream & out, goal const & g) { out << Z3_goal_to_string(g.ctx(), g); return out; }
|
||||||
};
|
};
|
||||||
|
|
||||||
|
@ -1283,8 +1306,7 @@ namespace z3 {
|
||||||
return *this;
|
return *this;
|
||||||
}
|
}
|
||||||
unsigned size() const { return Z3_apply_result_get_num_subgoals(ctx(), m_apply_result); }
|
unsigned size() const { return Z3_apply_result_get_num_subgoals(ctx(), m_apply_result); }
|
||||||
goal operator[](unsigned i) const { Z3_goal r = Z3_apply_result_get_subgoal(ctx(), m_apply_result, i); check_error(); return goal(ctx(), r); }
|
goal operator[](int i) const { assert(0 <= i); Z3_goal r = Z3_apply_result_get_subgoal(ctx(), m_apply_result, i); check_error(); return goal(ctx(), r); }
|
||||||
goal operator[](int i) const { assert(i >= 0); return this->operator[](static_cast<unsigned>(i)); }
|
|
||||||
model convert_model(model const & m, unsigned i = 0) const {
|
model convert_model(model const & m, unsigned i = 0) const {
|
||||||
check_context(*this, m);
|
check_context(*this, m);
|
||||||
Z3_model new_m = Z3_apply_result_convert_model(ctx(), m_apply_result, i, m);
|
Z3_model new_m = Z3_apply_result_convert_model(ctx(), m_apply_result, i, m);
|
||||||
|
@ -1437,6 +1459,17 @@ namespace z3 {
|
||||||
inline sort context::real_sort() { Z3_sort s = Z3_mk_real_sort(m_ctx); check_error(); return sort(*this, s); }
|
inline sort context::real_sort() { Z3_sort s = Z3_mk_real_sort(m_ctx); check_error(); return sort(*this, s); }
|
||||||
inline sort context::bv_sort(unsigned sz) { Z3_sort s = Z3_mk_bv_sort(m_ctx, sz); check_error(); return sort(*this, s); }
|
inline sort context::bv_sort(unsigned sz) { Z3_sort s = Z3_mk_bv_sort(m_ctx, sz); check_error(); return sort(*this, s); }
|
||||||
inline sort context::array_sort(sort d, sort r) { Z3_sort s = Z3_mk_array_sort(m_ctx, d, r); check_error(); return sort(*this, s); }
|
inline sort context::array_sort(sort d, sort r) { Z3_sort s = Z3_mk_array_sort(m_ctx, d, r); check_error(); return sort(*this, s); }
|
||||||
|
inline sort context::enumeration_sort(char const * name, unsigned n, char const * const * enum_names, func_decl_vector & cs, func_decl_vector & ts) {
|
||||||
|
array<Z3_symbol> _enum_names(n);
|
||||||
|
for (unsigned i = 0; i < n; i++) { _enum_names[i] = Z3_mk_string_symbol(*this, enum_names[i]); }
|
||||||
|
array<Z3_func_decl> _cs(n);
|
||||||
|
array<Z3_func_decl> _ts(n);
|
||||||
|
Z3_symbol _name = Z3_mk_string_symbol(*this, name);
|
||||||
|
sort s = to_sort(*this, Z3_mk_enumeration_sort(*this, _name, n, _enum_names.ptr(), _cs.ptr(), _ts.ptr()));
|
||||||
|
check_error();
|
||||||
|
for (unsigned i = 0; i < n; i++) { cs.push_back(func_decl(*this, _cs[i])); ts.push_back(func_decl(*this, _ts[i])); }
|
||||||
|
return s;
|
||||||
|
}
|
||||||
|
|
||||||
inline func_decl context::function(symbol const & name, unsigned arity, sort const * domain, sort const & range) {
|
inline func_decl context::function(symbol const & name, unsigned arity, sort const * domain, sort const & range) {
|
||||||
array<Z3_sort> args(arity);
|
array<Z3_sort> args(arity);
|
||||||
|
@ -1538,6 +1571,11 @@ namespace z3 {
|
||||||
return expr(ctx(), r);
|
return expr(ctx(), r);
|
||||||
|
|
||||||
}
|
}
|
||||||
|
inline expr func_decl::operator()() const {
|
||||||
|
Z3_ast r = Z3_mk_app(ctx(), *this, 0, 0);
|
||||||
|
ctx().check_error();
|
||||||
|
return expr(ctx(), r);
|
||||||
|
}
|
||||||
inline expr func_decl::operator()(expr const & a) const {
|
inline expr func_decl::operator()(expr const & a) const {
|
||||||
check_context(*this, a);
|
check_context(*this, a);
|
||||||
Z3_ast args[1] = { a };
|
Z3_ast args[1] = { a };
|
||||||
|
|
|
@ -78,6 +78,14 @@ namespace Microsoft.Z3
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
|
|
||||||
|
/// <summary>
|
||||||
|
/// Alias for Assert.
|
||||||
|
/// </summary>
|
||||||
|
public void Add(params BoolExpr[] constraints)
|
||||||
|
{
|
||||||
|
Assert(constraints);
|
||||||
|
}
|
||||||
|
|
||||||
/// <summary>
|
/// <summary>
|
||||||
/// Register predicate as recursive relation.
|
/// Register predicate as recursive relation.
|
||||||
/// </summary>
|
/// </summary>
|
||||||
|
|
|
@ -90,6 +90,14 @@ namespace Microsoft.Z3
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
|
|
||||||
|
/// <summary>
|
||||||
|
/// Alias for Assert.
|
||||||
|
/// </summary>
|
||||||
|
public void Add(params BoolExpr[] constraints)
|
||||||
|
{
|
||||||
|
Assert(constraints);
|
||||||
|
}
|
||||||
|
|
||||||
/// <summary>
|
/// <summary>
|
||||||
/// Indicates whether the goal contains `false'.
|
/// Indicates whether the goal contains `false'.
|
||||||
/// </summary>
|
/// </summary>
|
||||||
|
|
|
@ -117,6 +117,14 @@ namespace Microsoft.Z3
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
|
|
||||||
|
/// <summary>
|
||||||
|
/// Alias for Assert.
|
||||||
|
/// </summary>
|
||||||
|
public void Add(params BoolExpr[] constraints)
|
||||||
|
{
|
||||||
|
Assert(constraints);
|
||||||
|
}
|
||||||
|
|
||||||
/// <summary>
|
/// <summary>
|
||||||
/// Assert multiple constraints into the solver, and track them (in the unsat) core
|
/// Assert multiple constraints into the solver, and track them (in the unsat) core
|
||||||
/// using the Boolean constants in ps.
|
/// using the Boolean constants in ps.
|
||||||
|
|
|
@ -51,7 +51,7 @@ public class Fixedpoint extends Z3Object
|
||||||
*
|
*
|
||||||
* @throws Z3Exception
|
* @throws Z3Exception
|
||||||
**/
|
**/
|
||||||
public void assert_(BoolExpr ... constraints) throws Z3Exception
|
public void add(BoolExpr ... constraints) throws Z3Exception
|
||||||
{
|
{
|
||||||
getContext().checkContextMatch(constraints);
|
getContext().checkContextMatch(constraints);
|
||||||
for (BoolExpr a : constraints)
|
for (BoolExpr a : constraints)
|
||||||
|
|
|
@ -65,7 +65,7 @@ public class Goal extends Z3Object
|
||||||
*
|
*
|
||||||
* @throws Z3Exception
|
* @throws Z3Exception
|
||||||
**/
|
**/
|
||||||
public void assert_(BoolExpr ... constraints) throws Z3Exception
|
public void add(BoolExpr ... constraints) throws Z3Exception
|
||||||
{
|
{
|
||||||
getContext().checkContextMatch(constraints);
|
getContext().checkContextMatch(constraints);
|
||||||
for (BoolExpr c : constraints)
|
for (BoolExpr c : constraints)
|
||||||
|
|
|
@ -94,7 +94,7 @@ public class Solver extends Z3Object
|
||||||
*
|
*
|
||||||
* @throws Z3Exception
|
* @throws Z3Exception
|
||||||
**/
|
**/
|
||||||
public void assert_(BoolExpr... constraints) throws Z3Exception
|
public void add(BoolExpr... constraints) throws Z3Exception
|
||||||
{
|
{
|
||||||
getContext().checkContextMatch(constraints);
|
getContext().checkContextMatch(constraints);
|
||||||
for (BoolExpr a : constraints)
|
for (BoolExpr a : constraints)
|
||||||
|
|
|
@ -193,19 +193,9 @@ class hilbert_basis::value_index2 {
|
||||||
checker m_checker;
|
checker m_checker;
|
||||||
vector<numeral> m_keys;
|
vector<numeral> m_keys;
|
||||||
|
|
||||||
#if 1
|
|
||||||
numeral const* get_keys(values const& vs) {
|
numeral const* get_keys(values const& vs) {
|
||||||
return vs()-1;
|
return vs()-1;
|
||||||
}
|
}
|
||||||
#else
|
|
||||||
numeral const* get_keys(values const& vs) {
|
|
||||||
unsigned sz = m_keys.size();
|
|
||||||
for (unsigned i = 0; i < sz; ++i) {
|
|
||||||
m_keys[sz-i-1] = vs()[i-1];
|
|
||||||
}
|
|
||||||
return m_keys.c_ptr();
|
|
||||||
}
|
|
||||||
#endif
|
|
||||||
|
|
||||||
public:
|
public:
|
||||||
value_index2(hilbert_basis& hb): hb(hb), m_init(false) {
|
value_index2(hilbert_basis& hb): hb(hb), m_init(false) {
|
||||||
|
@ -507,7 +497,11 @@ class hilbert_basis::passive2 {
|
||||||
}
|
}
|
||||||
};
|
};
|
||||||
hilbert_basis& hb;
|
hilbert_basis& hb;
|
||||||
svector<offset_t> const& m_sos;
|
svector<offset_t> m_pos_sos;
|
||||||
|
svector<offset_t> m_neg_sos;
|
||||||
|
vector<numeral> m_pos_sos_sum;
|
||||||
|
vector<numeral> m_neg_sos_sum;
|
||||||
|
vector<numeral> m_sum_abs;
|
||||||
unsigned_vector m_psos;
|
unsigned_vector m_psos;
|
||||||
svector<offset_t> m_pas;
|
svector<offset_t> m_pas;
|
||||||
vector<numeral> m_weight;
|
vector<numeral> m_weight;
|
||||||
|
@ -528,40 +522,59 @@ class hilbert_basis::passive2 {
|
||||||
public:
|
public:
|
||||||
passive2(hilbert_basis& hb):
|
passive2(hilbert_basis& hb):
|
||||||
hb(hb),
|
hb(hb),
|
||||||
m_sos(hb.m_sos),
|
|
||||||
m_lt(&m_this),
|
m_lt(&m_this),
|
||||||
m_heap(10, m_lt)
|
m_heap(10, m_lt)
|
||||||
{
|
{
|
||||||
m_this = this;
|
m_this = this;
|
||||||
}
|
}
|
||||||
|
|
||||||
|
void init(svector<offset_t> const& I) {
|
||||||
|
for (unsigned i = 0; i < I.size(); ++i) {
|
||||||
|
numeral const& w = hb.vec(I[i]).weight();
|
||||||
|
if (w.is_pos()) {
|
||||||
|
m_pos_sos.push_back(I[i]);
|
||||||
|
m_pos_sos_sum.push_back(sum_abs(I[i]));
|
||||||
|
}
|
||||||
|
else {
|
||||||
|
m_neg_sos.push_back(I[i]);
|
||||||
|
m_neg_sos_sum.push_back(sum_abs(I[i]));
|
||||||
|
}
|
||||||
|
}
|
||||||
|
}
|
||||||
|
|
||||||
void reset() {
|
void reset() {
|
||||||
m_heap.reset();
|
m_heap.reset();
|
||||||
m_free_list.reset();
|
m_free_list.reset();
|
||||||
m_psos.reset();
|
m_psos.reset();
|
||||||
m_pas.reset();
|
m_pas.reset();
|
||||||
|
m_sum_abs.reset();
|
||||||
|
m_pos_sos.reset();
|
||||||
|
m_neg_sos.reset();
|
||||||
|
m_pos_sos_sum.reset();
|
||||||
|
m_neg_sos_sum.reset();
|
||||||
m_weight.reset();
|
m_weight.reset();
|
||||||
}
|
}
|
||||||
|
|
||||||
void insert(offset_t idx, unsigned offset) {
|
void insert(offset_t idx, unsigned offset) {
|
||||||
SASSERT(!m_sos.empty());
|
SASSERT(!m_pos_sos.empty());
|
||||||
unsigned v;
|
unsigned v;
|
||||||
numeral w = sum_abs(idx) + sum_abs(m_sos[0]);
|
|
||||||
if (m_free_list.empty()) {
|
if (m_free_list.empty()) {
|
||||||
v = m_pas.size();
|
v = m_pas.size();
|
||||||
m_pas.push_back(idx);
|
m_pas.push_back(idx);
|
||||||
m_psos.push_back(offset);
|
m_psos.push_back(offset);
|
||||||
m_weight.push_back(w);
|
m_weight.push_back(numeral(0));
|
||||||
m_heap.set_bounds(v+1);
|
m_heap.set_bounds(v+1);
|
||||||
|
m_sum_abs.push_back(sum_abs(idx));
|
||||||
}
|
}
|
||||||
else {
|
else {
|
||||||
v = m_free_list.back();
|
v = m_free_list.back();
|
||||||
m_free_list.pop_back();
|
m_free_list.pop_back();
|
||||||
m_pas[v] = idx;
|
m_pas[v] = idx;
|
||||||
m_psos[v] = offset;
|
m_psos[v] = offset;
|
||||||
m_weight[v] = w;
|
m_weight[v] = numeral(0);
|
||||||
|
m_sum_abs[v] = sum_abs(idx);
|
||||||
}
|
}
|
||||||
next_resolvable(v);
|
next_resolvable(hb.vec(idx).weight().is_pos(), v);
|
||||||
}
|
}
|
||||||
|
|
||||||
bool empty() const {
|
bool empty() const {
|
||||||
|
@ -571,12 +584,13 @@ public:
|
||||||
unsigned pop(offset_t& sos, offset_t& pas) {
|
unsigned pop(offset_t& sos, offset_t& pas) {
|
||||||
SASSERT (!empty());
|
SASSERT (!empty());
|
||||||
unsigned val = static_cast<unsigned>(m_heap.erase_min());
|
unsigned val = static_cast<unsigned>(m_heap.erase_min());
|
||||||
unsigned psos = m_psos[val];
|
|
||||||
sos = m_sos[psos];
|
|
||||||
pas = m_pas[val];
|
pas = m_pas[val];
|
||||||
m_psos[val]++;
|
|
||||||
next_resolvable(val);
|
|
||||||
numeral old_weight = hb.vec(pas).weight();
|
numeral old_weight = hb.vec(pas).weight();
|
||||||
|
bool is_positive = old_weight.is_pos();
|
||||||
|
unsigned psos = m_psos[val];
|
||||||
|
sos = is_positive?m_neg_sos[psos]:m_pos_sos[psos];
|
||||||
|
m_psos[val]++;
|
||||||
|
next_resolvable(is_positive, val);
|
||||||
numeral new_weight = hb.vec(sos).weight() + old_weight;
|
numeral new_weight = hb.vec(sos).weight() + old_weight;
|
||||||
if (new_weight.is_pos() != old_weight.is_pos()) {
|
if (new_weight.is_pos() != old_weight.is_pos()) {
|
||||||
psos = 0;
|
psos = 0;
|
||||||
|
@ -600,7 +614,7 @@ public:
|
||||||
public:
|
public:
|
||||||
iterator(passive2& p, unsigned i): p(p), m_idx(i) { fwd(); }
|
iterator(passive2& p, unsigned i): p(p), m_idx(i) { fwd(); }
|
||||||
offset_t pas() const { return p.m_pas[m_idx]; }
|
offset_t pas() const { return p.m_pas[m_idx]; }
|
||||||
offset_t sos() const { return p.m_sos[p.m_psos[m_idx]]; }
|
offset_t sos() const { return (p.hb.vec(pas()).weight().is_pos()?p.m_neg_sos:p.m_pos_sos)[p.m_psos[m_idx]]; }
|
||||||
iterator& operator++() { ++m_idx; fwd(); return *this; }
|
iterator& operator++() { ++m_idx; fwd(); return *this; }
|
||||||
iterator operator++(int) { iterator tmp = *this; ++*this; return tmp; }
|
iterator operator++(int) { iterator tmp = *this; ++*this; return tmp; }
|
||||||
bool operator==(iterator const& it) const {return m_idx == it.m_idx; }
|
bool operator==(iterator const& it) const {return m_idx == it.m_idx; }
|
||||||
|
@ -615,12 +629,14 @@ public:
|
||||||
return iterator(*this, m_pas.size());
|
return iterator(*this, m_pas.size());
|
||||||
}
|
}
|
||||||
private:
|
private:
|
||||||
void next_resolvable(unsigned v) {
|
void next_resolvable(bool is_positive, unsigned v) {
|
||||||
offset_t pas = m_pas[v];
|
offset_t pas = m_pas[v];
|
||||||
while (m_psos[v] < m_sos.size()) {
|
svector<offset_t> const& soss = is_positive?m_neg_sos:m_pos_sos;
|
||||||
offset_t sos = m_sos[m_psos[v]];
|
while (m_psos[v] < soss.size()) {
|
||||||
if (hb.can_resolve(sos, pas)) {
|
unsigned psos = m_psos[v];
|
||||||
m_weight[v] = sum_abs(pas) + sum_abs(sos);
|
offset_t sos = soss[psos];
|
||||||
|
if (hb.can_resolve(sos, pas, false)) {
|
||||||
|
m_weight[v] = m_sum_abs[v] + (is_positive?m_neg_sos_sum[psos]:m_pos_sos_sum[psos]);
|
||||||
m_heap.insert(v);
|
m_heap.insert(v);
|
||||||
return;
|
return;
|
||||||
}
|
}
|
||||||
|
@ -746,7 +762,7 @@ unsigned hilbert_basis::get_num_vars() const {
|
||||||
}
|
}
|
||||||
|
|
||||||
hilbert_basis::values hilbert_basis::vec(offset_t offs) const {
|
hilbert_basis::values hilbert_basis::vec(offset_t offs) const {
|
||||||
return values(m_store.c_ptr() + (get_num_vars() + 1)*offs.m_offset);
|
return values(m_ineqs.size(), m_store.c_ptr() + offs.m_offset);
|
||||||
}
|
}
|
||||||
|
|
||||||
void hilbert_basis::init_basis() {
|
void hilbert_basis::init_basis() {
|
||||||
|
@ -814,6 +830,9 @@ lbool hilbert_basis::saturate_orig(num_vector const& ineq, bool is_eq) {
|
||||||
offset_t idx = *it;
|
offset_t idx = *it;
|
||||||
values v = vec(idx);
|
values v = vec(idx);
|
||||||
v.weight() = get_weight(v, ineq);
|
v.weight() = get_weight(v, ineq);
|
||||||
|
for (unsigned k = 0; k < m_current_ineq; ++k) {
|
||||||
|
v.weight(k) = get_weight(v, m_ineqs[k]);
|
||||||
|
}
|
||||||
add_goal(idx);
|
add_goal(idx);
|
||||||
if (m_use_support) {
|
if (m_use_support) {
|
||||||
support.insert(idx.m_offset);
|
support.insert(idx.m_offset);
|
||||||
|
@ -833,7 +852,7 @@ lbool hilbert_basis::saturate_orig(num_vector const& ineq, bool is_eq) {
|
||||||
continue;
|
continue;
|
||||||
}
|
}
|
||||||
for (unsigned i = 0; !m_cancel && i < m_active.size(); ++i) {
|
for (unsigned i = 0; !m_cancel && i < m_active.size(); ++i) {
|
||||||
if ((!m_use_support || support.contains(m_active[i].m_offset)) && can_resolve(idx, m_active[i])) {
|
if ((!m_use_support || support.contains(m_active[i].m_offset)) && can_resolve(idx, m_active[i], true)) {
|
||||||
resolve(idx, m_active[i], j);
|
resolve(idx, m_active[i], j);
|
||||||
if (add_goal(j)) {
|
if (add_goal(j)) {
|
||||||
j = alloc_vector();
|
j = alloc_vector();
|
||||||
|
@ -884,6 +903,9 @@ lbool hilbert_basis::saturate(num_vector const& ineq, bool is_eq) {
|
||||||
offset_t idx = m_basis[i];
|
offset_t idx = m_basis[i];
|
||||||
values v = vec(idx);
|
values v = vec(idx);
|
||||||
v.weight() = get_weight(v, ineq);
|
v.weight() = get_weight(v, ineq);
|
||||||
|
for (unsigned k = 0; k < m_current_ineq; ++k) {
|
||||||
|
v.weight(k) = get_weight(v, m_ineqs[k]);
|
||||||
|
}
|
||||||
m_index->insert(idx, v);
|
m_index->insert(idx, v);
|
||||||
if (v.weight().is_zero()) {
|
if (v.weight().is_zero()) {
|
||||||
m_zero.push_back(idx);
|
m_zero.push_back(idx);
|
||||||
|
@ -896,6 +918,7 @@ lbool hilbert_basis::saturate(num_vector const& ineq, bool is_eq) {
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
m_basis.resize(init_basis_size);
|
m_basis.resize(init_basis_size);
|
||||||
|
m_passive2->init(m_sos);
|
||||||
// ASSERT basis is sorted by weight.
|
// ASSERT basis is sorted by weight.
|
||||||
|
|
||||||
// initialize passive
|
// initialize passive
|
||||||
|
@ -912,7 +935,7 @@ lbool hilbert_basis::saturate(num_vector const& ineq, bool is_eq) {
|
||||||
offset_t sos, pas;
|
offset_t sos, pas;
|
||||||
TRACE("hilbert_basis", display(tout); );
|
TRACE("hilbert_basis", display(tout); );
|
||||||
unsigned offset = m_passive2->pop(sos, pas);
|
unsigned offset = m_passive2->pop(sos, pas);
|
||||||
SASSERT(can_resolve(sos, pas));
|
SASSERT(can_resolve(sos, pas, true));
|
||||||
resolve(sos, pas, idx);
|
resolve(sos, pas, idx);
|
||||||
if (is_subsumed(idx)) {
|
if (is_subsumed(idx)) {
|
||||||
continue;
|
continue;
|
||||||
|
@ -1053,6 +1076,9 @@ void hilbert_basis::resolve(offset_t i, offset_t j, offset_t r) {
|
||||||
u[k] = v[k] + w[k];
|
u[k] = v[k] + w[k];
|
||||||
}
|
}
|
||||||
u.weight() = v.weight() + w.weight();
|
u.weight() = v.weight() + w.weight();
|
||||||
|
for (unsigned k = 0; k < m_current_ineq; ++k) {
|
||||||
|
u.weight(k) = v.weight(k) + w.weight(k);
|
||||||
|
}
|
||||||
TRACE("hilbert_basis_verbose",
|
TRACE("hilbert_basis_verbose",
|
||||||
display(tout, i);
|
display(tout, i);
|
||||||
display(tout, j);
|
display(tout, j);
|
||||||
|
@ -1063,10 +1089,11 @@ void hilbert_basis::resolve(offset_t i, offset_t j, offset_t r) {
|
||||||
|
|
||||||
hilbert_basis::offset_t hilbert_basis::alloc_vector() {
|
hilbert_basis::offset_t hilbert_basis::alloc_vector() {
|
||||||
if (m_free_list.empty()) {
|
if (m_free_list.empty()) {
|
||||||
unsigned num_vars = get_num_vars();
|
unsigned sz = m_ineqs.size() + get_num_vars();
|
||||||
unsigned idx = m_store.size();
|
unsigned idx = m_store.size();
|
||||||
m_store.resize(idx + 1 + num_vars);
|
m_store.resize(idx + sz);
|
||||||
return offset_t(idx/(1+num_vars));
|
// std::cout << "alloc vector: " << idx << " " << sz << " " << m_store.c_ptr() + idx << " " << m_ineqs.size() << "\n";
|
||||||
|
return offset_t(idx);
|
||||||
}
|
}
|
||||||
else {
|
else {
|
||||||
offset_t result = m_free_list.back();
|
offset_t result = m_free_list.back();
|
||||||
|
@ -1101,10 +1128,11 @@ bool hilbert_basis::is_subsumed(offset_t idx) {
|
||||||
return false;
|
return false;
|
||||||
}
|
}
|
||||||
|
|
||||||
bool hilbert_basis::can_resolve(offset_t i, offset_t j) const {
|
bool hilbert_basis::can_resolve(offset_t i, offset_t j, bool check_sign) const {
|
||||||
if (get_sign(i) == get_sign(j)) {
|
if (check_sign && get_sign(i) == get_sign(j)) {
|
||||||
return false;
|
return false;
|
||||||
}
|
}
|
||||||
|
SASSERT(get_sign(i) != get_sign(j));
|
||||||
values const& v1 = vec(i);
|
values const& v1 = vec(i);
|
||||||
values const& v2 = vec(j);
|
values const& v2 = vec(j);
|
||||||
if (v1[0].is_one() && v2[0].is_one()) {
|
if (v1[0].is_one() && v2[0].is_one()) {
|
||||||
|
@ -1123,7 +1151,7 @@ bool hilbert_basis::can_resolve(offset_t i, offset_t j) const {
|
||||||
}
|
}
|
||||||
|
|
||||||
hilbert_basis::sign_t hilbert_basis::get_sign(offset_t idx) const {
|
hilbert_basis::sign_t hilbert_basis::get_sign(offset_t idx) const {
|
||||||
numeral val = vec(idx).weight();
|
numeral const& val = vec(idx).weight();
|
||||||
if (val.is_pos()) {
|
if (val.is_pos()) {
|
||||||
return pos;
|
return pos;
|
||||||
}
|
}
|
||||||
|
@ -1267,7 +1295,7 @@ bool hilbert_basis::is_subsumed(offset_t i, offset_t j) const {
|
||||||
n >= m && (!m.is_neg() || n == m) &&
|
n >= m && (!m.is_neg() || n == m) &&
|
||||||
is_geq(v, w);
|
is_geq(v, w);
|
||||||
for (unsigned k = 0; r && k < m_current_ineq; ++k) {
|
for (unsigned k = 0; r && k < m_current_ineq; ++k) {
|
||||||
r = get_weight(vec(i), m_ineqs[k]) >= get_weight(vec(j), m_ineqs[k]);
|
r = v.weight(k) >= w.weight(k);
|
||||||
}
|
}
|
||||||
CTRACE("hilbert_basis", r,
|
CTRACE("hilbert_basis", r,
|
||||||
display(tout, i);
|
display(tout, i);
|
||||||
|
|
|
@ -56,12 +56,14 @@ private:
|
||||||
class values {
|
class values {
|
||||||
numeral* m_values;
|
numeral* m_values;
|
||||||
public:
|
public:
|
||||||
values(numeral* v):m_values(v) {}
|
values(unsigned offset, numeral* v): m_values(v+offset) { }
|
||||||
numeral& weight() { return m_values[0]; } // value of a*x
|
numeral& weight() { return m_values[-1]; } // value of a*x
|
||||||
numeral& operator[](unsigned i) { return m_values[i+1]; } // value of x_i
|
numeral const& weight() const { return m_values[-1]; } // value of a*x
|
||||||
numeral const& weight() const { return m_values[0]; } // value of a*x
|
numeral& weight(int i) { return m_values[-2-i]; } // value of b_i*x for 0 <= i < current inequality.
|
||||||
numeral const& operator[](unsigned i) const { return m_values[i+1]; } // value of x_i
|
numeral const& weight(int i) const { return m_values[-2-i]; } // value of b_i*x
|
||||||
numeral const* operator()() const { return m_values + 1; }
|
numeral& operator[](unsigned i) { return m_values[i]; } // value of x_i
|
||||||
|
numeral const& operator[](unsigned i) const { return m_values[i]; } // value of x_i
|
||||||
|
numeral const* operator()() const { return m_values; }
|
||||||
};
|
};
|
||||||
|
|
||||||
vector<num_vector> m_ineqs; // set of asserted inequalities
|
vector<num_vector> m_ineqs; // set of asserted inequalities
|
||||||
|
@ -114,7 +116,7 @@ private:
|
||||||
bool is_subsumed(offset_t idx);
|
bool is_subsumed(offset_t idx);
|
||||||
bool is_subsumed(offset_t i, offset_t j) const;
|
bool is_subsumed(offset_t i, offset_t j) const;
|
||||||
void recycle(offset_t idx);
|
void recycle(offset_t idx);
|
||||||
bool can_resolve(offset_t i, offset_t j) const;
|
bool can_resolve(offset_t i, offset_t j, bool check_sign) const;
|
||||||
sign_t get_sign(offset_t idx) const;
|
sign_t get_sign(offset_t idx) const;
|
||||||
bool add_goal(offset_t idx);
|
bool add_goal(offset_t idx);
|
||||||
offset_t alloc_vector();
|
offset_t alloc_vector();
|
||||||
|
|
|
@ -521,6 +521,8 @@ void tst_hilbert_basis() {
|
||||||
tst2();
|
tst2();
|
||||||
tst3();
|
tst3();
|
||||||
tst4();
|
tst4();
|
||||||
|
tst4();
|
||||||
|
tst4();
|
||||||
tst5();
|
tst5();
|
||||||
tst6();
|
tst6();
|
||||||
tst7();
|
tst7();
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue