SyGuS: Functions
examples/api/cpp/sygus-fun.cpp
  1/******************************************************************************
  2 * Top contributors (to current version):
  3 *   Abdalrhman Mohamed, Aina Niemetz, Mathias Preiner
  4 *
  5 * This file is part of the cvc5 project.
  6 *
  7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
  8 * in the top-level source directory and their institutional affiliations.
  9 * All rights reserved.  See the file COPYING in the top-level source
 10 * directory for licensing information.
 11 * ****************************************************************************
 12 *
 13 * A simple demonstration of the Sygus API.
 14 *
 15 * A simple demonstration of how to use the Sygus API to synthesize max and min
 16 * functions.
 17 */
 18
 19#include <cvc5/cvc5.h>
 20
 21#include <iostream>
 22
 23#include "utils.h"
 24
 25using namespace cvc5;
 26
 27int main()
 28{
 29  TermManager tm;
 30  Solver slv(tm);
 31
 32  // required options
 33  slv.setOption("sygus", "true");
 34  slv.setOption("incremental", "false");
 35
 36  // set the logic
 37  slv.setLogic("LIA");
 38
 39  Sort integer = tm.getIntegerSort();
 40  Sort boolean = tm.getBooleanSort();
 41
 42  // declare input variables for the functions-to-synthesize
 43  Term x = tm.mkVar(integer, "x");
 44  Term y = tm.mkVar(integer, "y");
 45
 46  // declare the grammar non-terminals
 47  Term start = tm.mkVar(integer, "Start");
 48  Term start_bool = tm.mkVar(boolean, "StartBool");
 49
 50  // define the rules
 51  Term zero = tm.mkInteger(0);
 52  Term one = tm.mkInteger(1);
 53
 54  Term plus = tm.mkTerm(Kind::ADD, {start, start});
 55  Term minus = tm.mkTerm(Kind::SUB, {start, start});
 56  Term ite = tm.mkTerm(Kind::ITE, {start_bool, start, start});
 57
 58  Term And = tm.mkTerm(Kind::AND, {start_bool, start_bool});
 59  Term Not = tm.mkTerm(Kind::NOT, {start_bool});
 60  Term leq = tm.mkTerm(Kind::LEQ, {start, start});
 61
 62  // create the grammar object
 63  Grammar g = slv.mkGrammar({x, y}, {start, start_bool});
 64
 65  // bind each non-terminal to its rules
 66  g.addRules(start, {zero, one, x, y, plus, minus, ite});
 67  g.addRules(start_bool, {And, Not, leq});
 68
 69  // declare the functions-to-synthesize. Optionally, provide the grammar
 70  // constraints
 71  Term max = slv.synthFun("max", {x, y}, integer, g);
 72  Term min = slv.synthFun("min", {x, y}, integer);
 73
 74  // declare universal variables.
 75  Term varX = slv.declareSygusVar("x", integer);
 76  Term varY = slv.declareSygusVar("y", integer);
 77
 78  Term max_x_y = tm.mkTerm(Kind::APPLY_UF, {max, varX, varY});
 79  Term min_x_y = tm.mkTerm(Kind::APPLY_UF, {min, varX, varY});
 80
 81  // add semantic constraints
 82  // (constraint (>= (max x y) x))
 83  slv.addSygusConstraint(tm.mkTerm(Kind::GEQ, {max_x_y, varX}));
 84
 85  // (constraint (>= (max x y) y))
 86  slv.addSygusConstraint(tm.mkTerm(Kind::GEQ, {max_x_y, varY}));
 87
 88  // (constraint (or (= x (max x y))
 89  //                 (= y (max x y))))
 90  slv.addSygusConstraint(tm.mkTerm(Kind::OR,
 91                                   {tm.mkTerm(Kind::EQUAL, {max_x_y, varX}),
 92                                    tm.mkTerm(Kind::EQUAL, {max_x_y, varY})}));
 93
 94  // (constraint (= (+ (max x y) (min x y))
 95  //                (+ x y)))
 96  slv.addSygusConstraint(tm.mkTerm(Kind::EQUAL,
 97                                   {tm.mkTerm(Kind::ADD, {max_x_y, min_x_y}),
 98                                    tm.mkTerm(Kind::ADD, {varX, varY})}));
 99
100  // print solutions if available
101  if (slv.checkSynth().hasSolution())
102  {
103    // Output should be equivalent to:
104    // (
105    //   (define-fun max ((x Int) (y Int)) Int (ite (<= x y) y x))
106    //   (define-fun min ((x Int) (y Int)) Int (ite (<= x y) x y))
107    // )
108    std::vector<Term> terms = {max, min};
109    utils::printSynthSolutions(terms, slv.getSynthSolutions(terms));
110  }
111
112  return 0;
113}
  1/******************************************************************************
  2 * Top contributors (to current version):
  3 *   Aina Niemetz
  4 *
  5 * This file is part of the cvc5 project.
  6 *
  7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
  8 * in the top-level source directory and their institutional affiliations.
  9 * All rights reserved.  See the file COPYING in the top-level source
 10 * directory for licensing information.
 11 * ****************************************************************************
 12 *
 13 * A simple demonstration of the Sygus API.
 14 *
 15 * A simple demonstration of how to use the Sygus API to synthesize max and min
 16 * functions.
 17 */
 18
 19#include <cvc5/c/cvc5.h>
 20#include <stdio.h>
 21
 22#include "utils.h"
 23
 24int main()
 25{
 26  Cvc5TermManager* tm = cvc5_term_manager_new();
 27  Cvc5* slv = cvc5_new(tm);
 28
 29  // required options
 30  cvc5_set_option(slv, "sygus", "true");
 31  cvc5_set_option(slv, "incremental", "false");
 32
 33  // set the logic
 34  cvc5_set_logic(slv, "LIA");
 35
 36  Cvc5Sort int_sort = cvc5_get_integer_sort(tm);
 37  Cvc5Sort bool_sort = cvc5_get_boolean_sort(tm);
 38
 39  // declare input variables for the functions-to-synthesize
 40  Cvc5Term x = cvc5_mk_var(tm, int_sort, "x");
 41  Cvc5Term y = cvc5_mk_var(tm, int_sort, "y");
 42
 43  // declare the grammar non-terminals
 44  Cvc5Term start = cvc5_mk_var(tm, int_sort, "Start");
 45  Cvc5Term start_bool = cvc5_mk_var(tm, bool_sort, "StartBool");
 46
 47  // define the rules
 48  Cvc5Term zero = cvc5_mk_integer_int64(tm, 0);
 49  Cvc5Term one = cvc5_mk_integer_int64(tm, 1);
 50
 51  Cvc5Term args2[2] = {start, start};
 52  Cvc5Term plus = cvc5_mk_term(tm, CVC5_KIND_ADD, 2, args2);
 53  Cvc5Term minus = cvc5_mk_term(tm, CVC5_KIND_SUB, 2, args2);
 54  Cvc5Term args3[3] = {start_bool, start, start};
 55  Cvc5Term ite = cvc5_mk_term(tm, CVC5_KIND_ITE, 3, args3);
 56
 57  args2[0] = start_bool;
 58  args2[1] = start_bool;
 59  Cvc5Term And = cvc5_mk_term(tm, CVC5_KIND_AND, 2, args2);
 60  Cvc5Term args1[1] = {start_bool};
 61  Cvc5Term Not = cvc5_mk_term(tm, CVC5_KIND_NOT, 1, args1);
 62  args2[0] = start;
 63  args2[1] = start;
 64  Cvc5Term leq = cvc5_mk_term(tm, CVC5_KIND_LEQ, 2, args2);
 65
 66  // create the grammar object
 67  Cvc5Term b[2] = {x, y};
 68  Cvc5Term s[2] = {start, start_bool};
 69  Cvc5Grammar g = cvc5_mk_grammar(slv, 2, b, 2, s);
 70
 71  // bind each non-terminal to its rules
 72  Cvc5Term rules7[7] = {zero, one, x, y, plus, minus, ite};
 73  cvc5_grammar_add_rules(g, start, 7, rules7);
 74  Cvc5Term rules3[3] = {And, Not, leq};
 75  cvc5_grammar_add_rules(g, start_bool, 3, rules3);
 76
 77  // declare the functions-to-synthesize. Optionally, provide the grammar
 78  // constraints
 79  Cvc5Term max = cvc5_synth_fun_with_grammar(slv, "max", 2, b, int_sort, g);
 80  Cvc5Term min = cvc5_synth_fun(slv, "min", 2, b, int_sort);
 81
 82  // declare universal variables.
 83  Cvc5Term var_x = cvc5_declare_sygus_var(slv, "x", int_sort);
 84  Cvc5Term var_y = cvc5_declare_sygus_var(slv, "y", int_sort);
 85
 86  args3[0] = max;
 87  args3[1] = var_x;
 88  args3[2] = var_y;
 89  Cvc5Term max_x_y = cvc5_mk_term(tm, CVC5_KIND_APPLY_UF, 3, args3);
 90  args3[0] = min;
 91  Cvc5Term min_x_y = cvc5_mk_term(tm, CVC5_KIND_APPLY_UF, 3, args3);
 92
 93  // add semantic constraints
 94  // (constraint (>= (max x y) x))
 95  args2[0] = max_x_y;
 96  args2[1] = var_x;
 97  cvc5_add_sygus_constraint(slv, cvc5_mk_term(tm, CVC5_KIND_GEQ, 2, args2));
 98
 99  // (constraint (>= (max x y) y))
100  args2[1] = var_y;
101  cvc5_add_sygus_constraint(slv, cvc5_mk_term(tm, CVC5_KIND_GEQ, 2, args2));
102
103  // (constraint (or (= x (max x y))
104  //                 (= y (max x y))))
105  args2[0] = max_x_y;
106  args2[1] = var_x;
107  Cvc5Term lhs = cvc5_mk_term(tm, CVC5_KIND_EQUAL, 2, args2);
108  args2[0] = max_x_y;
109  args2[1] = var_y;
110  Cvc5Term rhs = cvc5_mk_term(tm, CVC5_KIND_EQUAL, 2, args2);
111  args2[0] = lhs;
112  args2[1] = rhs;
113  cvc5_add_sygus_constraint(slv, cvc5_mk_term(tm, CVC5_KIND_OR, 2, args2));
114
115  // (constraint (= (+ (max x y) (min x y))
116  //                (+ x y)))
117  args2[0] = max_x_y;
118  args2[1] = min_x_y;
119  cvc5_term_release(lhs);  // optional, not needed anymore so we can release
120  lhs = cvc5_mk_term(tm, CVC5_KIND_ADD, 2, args2);
121  args2[0] = var_x;
122  args2[1] = var_y;
123  cvc5_term_release(rhs);  // optional, not needed anymore so we can release
124  rhs = cvc5_mk_term(tm, CVC5_KIND_ADD, 2, args2);
125  args2[0] = lhs;
126  args2[1] = rhs;
127  cvc5_add_sygus_constraint(slv, cvc5_mk_term(tm, CVC5_KIND_EQUAL, 2, args2));
128
129  // print solutions if available
130  if (cvc5_synth_result_has_solution(cvc5_check_synth(slv)))
131  {
132    // Output should be equivalent to:
133    // (
134    //   (define-fun max ((x Int) (y Int)) Int (ite (<= x y) y x))
135    //   (define-fun min ((x Int) (y Int)) Int (ite (<= x y) x y))
136    // )
137    args2[0] = max;
138    args2[1] = min;
139    const Cvc5Term* sols = cvc5_get_synth_solutions(slv, 2, args2);
140    print_synth_solutions(2, args2, sols);
141  }
142
143  cvc5_delete(slv);
144  cvc5_term_manager_delete(tm);
145  return 0;
146}
examples/api/java/SygusFun.java
  1/******************************************************************************
  2 * Top contributors (to current version):
  3 *   Mudathir Mohamed, Daniel Larraz, Mathias Preiner
  4 *
  5 * This file is part of the cvc5 project.
  6 *
  7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
  8 * in the top-level source directory and their institutional affiliations.
  9 * All rights reserved.  See the file COPYING in the top-level source
 10 * directory for licensing information.
 11 * ****************************************************************************
 12 *
 13 * A simple demonstration of the Sygus API.
 14 *
 15 * A simple demonstration of how to use the Sygus API to synthesize max and min
 16 * functions. This is a direct translation of sygus-fun.cpp.
 17 */
 18
 19import static io.github.cvc5.Kind.*;
 20
 21import io.github.cvc5.*;
 22
 23public class SygusFun
 24{
 25  public static void main(String args[]) throws CVC5ApiException
 26  {
 27    TermManager tm = new TermManager();
 28    Solver slv = new Solver(tm);
 29    {
 30      // required options
 31      slv.setOption("sygus", "true");
 32      slv.setOption("incremental", "false");
 33
 34      // set the logic
 35      slv.setLogic("LIA");
 36
 37      Sort integer = tm.getIntegerSort();
 38      Sort bool = tm.getBooleanSort();
 39
 40      // declare input variables for the functions-to-synthesize
 41      Term x = tm.mkVar(integer, "x");
 42      Term y = tm.mkVar(integer, "y");
 43
 44      // declare the grammar non-terminals
 45      Term start = tm.mkVar(integer, "Start");
 46      Term start_bool = tm.mkVar(bool, "StartBool");
 47
 48      // define the rules
 49      Term zero = tm.mkInteger(0);
 50      Term one = tm.mkInteger(1);
 51
 52      Term plus = tm.mkTerm(ADD, start, start);
 53      Term minus = tm.mkTerm(SUB, start, start);
 54      Term ite = tm.mkTerm(ITE, start_bool, start, start);
 55
 56      Term And = tm.mkTerm(AND, start_bool, start_bool);
 57      Term Not = tm.mkTerm(NOT, start_bool);
 58      Term leq = tm.mkTerm(LEQ, start, start);
 59
 60      // create the grammar object
 61      Grammar g = slv.mkGrammar(new Term[] {x, y}, new Term[] {start, start_bool});
 62
 63      // bind each non-terminal to its rules
 64      g.addRules(start, new Term[] {zero, one, x, y, plus, minus, ite});
 65      g.addRules(start_bool, new Term[] {And, Not, leq});
 66
 67      // declare the functions-to-synthesize. Optionally, provide the grammar
 68      // constraints
 69      Term max = slv.synthFun("max", new Term[] {x, y}, integer, g);
 70      Term min = slv.synthFun("min", new Term[] {x, y}, integer);
 71
 72      // declare universal variables.
 73      Term varX = slv.declareSygusVar("x", integer);
 74      Term varY = slv.declareSygusVar("y", integer);
 75
 76      Term max_x_y = tm.mkTerm(APPLY_UF, max, varX, varY);
 77      Term min_x_y = tm.mkTerm(APPLY_UF, min, varX, varY);
 78
 79      // add semantic constraints
 80      // (constraint (>= (max x y) x))
 81      slv.addSygusConstraint(tm.mkTerm(GEQ, max_x_y, varX));
 82
 83      // (constraint (>= (max x y) y))
 84      slv.addSygusConstraint(tm.mkTerm(GEQ, max_x_y, varY));
 85
 86      // (constraint (or (= x (max x y))
 87      //                 (= y (max x y))))
 88      slv.addSygusConstraint(
 89          tm.mkTerm(OR, tm.mkTerm(EQUAL, max_x_y, varX), tm.mkTerm(EQUAL, max_x_y, varY)));
 90
 91      // (constraint (= (+ (max x y) (min x y))
 92      //                (+ x y)))
 93      slv.addSygusConstraint(
 94          tm.mkTerm(EQUAL, tm.mkTerm(ADD, max_x_y, min_x_y), tm.mkTerm(ADD, varX, varY)));
 95
 96      // print solutions if available
 97      if (slv.checkSynth().hasSolution())
 98      {
 99        // Output should be equivalent to:
100        // (
101        //   (define-fun max ((x Int) (y Int)) Int (ite (<= x y) y x))
102        //   (define-fun min ((x Int) (y Int)) Int (ite (<= x y) x y))
103        // )
104        Term[] terms = new Term[] {max, min};
105        Utils.printSynthSolutions(terms, slv.getSynthSolutions(terms));
106      }
107    }
108    Context.deletePointers();
109  }
110}
examples/api/python/sygus-fun.py
  1#!/usr/bin/env python
  2###############################################################################
  3# Top contributors (to current version):
  4#   Yoni Zohar, Aina Niemetz, Andrew Reynolds
  5#
  6# This file is part of the cvc5 project.
  7#
  8# Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
  9# in the top-level source directory and their institutional affiliations.
 10# All rights reserved.  See the file COPYING in the top-level source
 11# directory for licensing information.
 12# #############################################################################
 13#
 14# A simple demonstration of the solving capabilities of the cvc5 sygus solver
 15# through the Python API. This is a direct translation of sygus-fun.cpp.
 16##
 17
 18import copy
 19import cvc5
 20import utils
 21from cvc5 import Kind
 22
 23if __name__ == "__main__":
 24  tm = cvc5.TermManager()
 25  slv = cvc5.Solver(tm)
 26
 27  # required options
 28  slv.setOption("sygus", "true")
 29  slv.setOption("incremental", "false")
 30
 31  # set the logic
 32  slv.setLogic("LIA")
 33
 34  integer = tm.getIntegerSort()
 35  boolean = tm.getBooleanSort()
 36
 37  # declare input variables for the functions-to-synthesize
 38  x = tm.mkVar(integer, "x")
 39  y = tm.mkVar(integer, "y")
 40
 41  # declare the grammar non-terminals
 42  start = tm.mkVar(integer, "Start")
 43  start_bool = tm.mkVar(boolean, "StartBool")
 44
 45  # define the rules
 46  zero = tm.mkInteger(0)
 47  one = tm.mkInteger(1)
 48
 49  plus = tm.mkTerm(Kind.ADD, start, start)
 50  minus = tm.mkTerm(Kind.SUB, start, start)
 51  ite = tm.mkTerm(Kind.ITE, start_bool, start, start)
 52
 53  And = tm.mkTerm(Kind.AND, start_bool, start_bool)
 54  Not = tm.mkTerm(Kind.NOT, start_bool)
 55  leq = tm.mkTerm(Kind.LEQ, start, start)
 56
 57  # create the grammar object
 58  g = slv.mkGrammar([x, y], [start, start_bool])
 59
 60  # bind each non-terminal to its rules
 61  g.addRules(start, [zero, one, x, y, plus, minus, ite])
 62  g.addRules(start_bool, [And, Not, leq])
 63
 64  # declare the functions-to-synthesize. Optionally, provide the grammar
 65  # constraints
 66  max = slv.synthFun("max", [x, y], integer, g)
 67  min = slv.synthFun("min", [x, y], integer)
 68
 69  # declare universal variables.
 70  varX = slv.declareSygusVar("x", integer)
 71  varY = slv.declareSygusVar("y", integer)
 72
 73  max_x_y = tm.mkTerm(Kind.APPLY_UF, max, varX, varY)
 74  min_x_y = tm.mkTerm(Kind.APPLY_UF, min, varX, varY)
 75
 76  # add semantic constraints
 77  # (constraint (>= (max x y) x))
 78  slv.addSygusConstraint(tm.mkTerm(Kind.GEQ, max_x_y, varX))
 79
 80  # (constraint (>= (max x y) y))
 81  slv.addSygusConstraint(tm.mkTerm(Kind.GEQ, max_x_y, varY))
 82
 83  # (constraint (or (= x (max x y))
 84  #                 (= y (max x y))))
 85  slv.addSygusConstraint(tm.mkTerm(
 86      Kind.OR,
 87      tm.mkTerm(Kind.EQUAL, max_x_y, varX),
 88      tm.mkTerm(Kind.EQUAL, max_x_y, varY)))
 89
 90  # (constraint (= (+ (max x y) (min x y))
 91  #                (+ x y)))
 92  slv.addSygusConstraint(tm.mkTerm(
 93      Kind.EQUAL,
 94      tm.mkTerm(Kind.ADD, max_x_y, min_x_y),
 95      tm.mkTerm(Kind.ADD, varX, varY)))
 96
 97  # print solutions if available
 98  if (slv.checkSynth().hasSolution()):
 99    # Output should be equivalent to:
100    # (define-fun max ((x Int) (y Int)) Int (ite (<= x y) y x))
101    # (define-fun min ((x Int) (y Int)) Int (ite (<= x y) x y))
102    terms = [max, min]
103    utils.print_synth_solutions(terms, slv.getSynthSolutions(terms))
examples/api/smtlib/sygus-fun.sy
 1; The printed output for this example should be equivalent to:
 2; (
 3;   (define-fun max ((x Int) (y Int)) Int (ite (<= x y) y x))
 4;   (define-fun min ((x Int) (y Int)) Int (ite (<= x y) x y))
 5; )
 6
 7(set-logic LIA)
 8(synth-fun max ((x Int) (y Int)) Int
 9  ((Start Int) (StartBool Bool))
10  ((Start Int (0 1 x y
11               (+ Start Start)
12               (- Start Start)
13               (ite StartBool Start Start)))
14   (StartBool Bool ((and StartBool StartBool)
15                    (not StartBool)
16                    (<= Start Start)))))
17(synth-fun min ((x Int) (y Int)) Int)
18(declare-var x Int)
19(declare-var y Int)
20(constraint (>= (max x y) x))
21(constraint (>= (max x y) y))
22(constraint (or (= x (max x y)) (= y (max x y))))
23(constraint (= (+ (max x y) (min x y)) (+ x y)))
24(check-synth)
The utility method used for printing the synthesis solutions in the examples
above is defined separately in the utils module:
 1/******************************************************************************
 2 * Top contributors (to current version):
 3 *   Abdalrhman Mohamed, Mathias Preiner
 4 *
 5 * This file is part of the cvc5 project.
 6 *
 7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
 8 * in the top-level source directory and their institutional affiliations.
 9 * All rights reserved.  See the file COPYING in the top-level source
10 * directory for licensing information.
11 * ****************************************************************************
12 *
13 * Utility methods.
14 */
15
16#ifndef CVC5__UTILS_H
17#define CVC5__UTILS_H
18
19#include <cvc5/cvc5.h>
20
21namespace utils {
22
23/**
24 * Print solutions for synthesis conjecture to the standard output stream.
25 * @param terms the terms for which the synthesis solutions were retrieved
26 * @param sols the synthesis solutions of the given terms
27 */
28void printSynthSolutions(const std::vector<cvc5::Term>& terms,
29                         const std::vector<cvc5::Term>& sols);
30
31}  // namespace utils
32
33#endif  // CVC5__UTILS_H
 1/******************************************************************************
 2 * Top contributors (to current version):
 3 *   Abdalrhman Mohamed, Mathias Preiner, Yoni Zohar
 4 *
 5 * This file is part of the cvc5 project.
 6 *
 7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
 8 * in the top-level source directory and their institutional affiliations.
 9 * All rights reserved.  See the file COPYING in the top-level source
10 * directory for licensing information.
11 * ****************************************************************************
12 *
13 * Implementations of utility methods.
14 */
15
16#include "utils.h"
17
18#include <iostream>
19
20namespace utils {
21
22using namespace cvc5;
23
24/**
25 * Get the string version of define-fun command.
26 * @param f the function to print
27 * @param params the function parameters
28 * @param body the function body
29 * @return a string version of define-fun
30 */
31std::string defineFunToString(const cvc5::Term& f,
32                              const std::vector<cvc5::Term>& params,
33                              const cvc5::Term& body)
34{
35  cvc5::Sort sort = f.getSort();
36  if (sort.isFunction())
37  {
38    sort = sort.getFunctionCodomainSort();
39  }
40  std::stringstream ss;
41  ss << "(define-fun " << f << " (";
42  for (size_t i = 0, n = params.size(); i < n; ++i)
43  {
44    if (i > 0)
45    {
46      ss << ' ';
47    }
48    ss << '(' << params[i] << ' ' << params[i].getSort() << ')';
49  }
50  ss << ") " << sort << ' ' << body << ')';
51  return ss.str();
52}
53
54void printSynthSolutions(const std::vector<cvc5::Term>& terms,
55                         const std::vector<cvc5::Term>& sols)
56{
57  std::cout << '(' << std::endl;
58  for (size_t i = 0, n = terms.size(); i < n; ++i)
59  {
60    std::vector<cvc5::Term> params;
61    cvc5::Term body = sols[i];
62    if (sols[i].getKind() == cvc5::Kind::LAMBDA)
63    {
64      params.insert(params.end(), sols[i][0].begin(), sols[i][0].end());
65      body = sols[i][1];
66    }
67    std::cout << "  " << defineFunToString(terms[i], params, body) << std::endl;
68  }
69  std::cout << ')' << std::endl;
70}
71
72}  // namespace utils
 1/******************************************************************************
 2 * Top contributors (to current version):
 3 *   Aina Niemetz
 4 *
 5 * This file is part of the cvc5 project.
 6 *
 7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
 8 * in the top-level source directory and their institutional affiliations.
 9 * All rights reserved.  See the file COPYING in the top-level source
10 * directory for licensing information.
11 * ****************************************************************************
12 *
13 * Utility methods.
14 */
15
16#ifndef CVC5__C_UTILS_H
17#define CVC5__C_UTILS_H
18
19#include <cvc5/c/cvc5.h>
20
21/**
22 * Print solutions for synthesis conjecture to the stdout.
23 * @param nterms The number of terms.
24 * @param terms  The terms for which the synthesis solutions were retrieved.
25 * @param sols   The synthesis solutions of the given terms.
26 */
27void print_synth_solutions(size_t nterms,
28                           const Cvc5Term terms[],
29                           const Cvc5Term sols[]);
30
31#endif
 1/******************************************************************************
 2 * Top contributors (to current version):
 3 *   Aina Niemetz
 4 *
 5 * This file is part of the cvc5 project.
 6 *
 7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
 8 * in the top-level source directory and their institutional affiliations.
 9 * All rights reserved.  See the file COPYING in the top-level source
10 * directory for licensing information.
11 * ****************************************************************************
12 *
13 * Implementations of utility methods.
14 */
15
16#include "utils.h"
17
18#include <stdio.h>
19#include <stdlib.h>
20
21/**
22 * Get the string version of define-fun command.
23 * @param f       The function to print.
24 * @param nparams The number of function parameters.
25 * @param params  The function parameters.
26 * @param body    The function body.
27 * @return A string version of define-fun.
28 */
29void print_define_fun(const Cvc5Term f,
30                      size_t nparams,
31                      const Cvc5Term params[],
32                      const Cvc5Term body)
33{
34  Cvc5Sort sort = cvc5_term_get_sort(f);
35  if (cvc5_sort_is_fun(sort))
36  {
37    sort = cvc5_sort_fun_get_codomain(sort);
38  }
39  printf("(define-fun %s (", cvc5_term_to_string(f));
40  for (size_t i = 0; i < nparams; i++)
41  {
42    printf("%s", i > 0 ? " " : "");
43    printf("(%s %s)",
44           cvc5_term_to_string(params[i]),
45           cvc5_sort_to_string(cvc5_term_get_sort(params[i])));
46  }
47  printf(") %s %s)", cvc5_sort_to_string(sort), cvc5_term_to_string(body));
48}
49
50void print_synth_solutions(size_t nterms,
51                           const Cvc5Term terms[],
52                           const Cvc5Term sols[])
53{
54  printf("(\n");
55  for (size_t i = 0; i < nterms; i++)
56  {
57    size_t nparams = 0;
58    Cvc5Term* params = NULL;
59    Cvc5Term body = sols[i];
60    if (cvc5_term_get_kind(sols[i]) == CVC5_KIND_LAMBDA)
61    {
62      Cvc5Term psols = cvc5_term_get_child(sols[i], 0);
63      nparams = cvc5_term_get_num_children(psols);
64      params = (Cvc5Term*)malloc(nparams * sizeof(Cvc5Term));
65      for (size_t k = 0; k < nparams; k++)
66      {
67        params[k] = cvc5_term_get_child(psols, k);
68      }
69      body = cvc5_term_get_child(sols[i], 1);
70    }
71    printf("  ");
72    print_define_fun(terms[i], nterms, params, body);
73    if (params)
74    {
75      free(params);
76    }
77    printf("\n");
78  }
79  printf(")\n");
80}
 1/******************************************************************************
 2 * Top contributors (to current version):
 3 *   Mudathir Mohamed, Abdalrhman Mohamed, Andres Noetzli
 4 *
 5 * This file is part of the cvc5 project.
 6 *
 7 * Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
 8 * in the top-level source directory and their institutional affiliations.
 9 * All rights reserved.  See the file COPYING in the top-level source
10 * directory for licensing information.
11 * ****************************************************************************
12 *
13 * Utility methods.
14 */
15
16import io.github.cvc5.CVC5ApiException;
17import io.github.cvc5.Kind;
18import io.github.cvc5.Sort;
19import io.github.cvc5.Term;
20import java.util.ArrayList;
21import java.util.List;
22
23public class Utils
24{
25  /**
26   * Get the string version of define-fun command.
27   *
28   * @param f the function to print
29   * @param params the function parameters
30   * @param body the function body
31   * @return a string version of define-fun
32   */
33  private static String defineFunToString(Term f, Term[] params, Term body)
34  {
35    Sort sort = f.getSort();
36    if (sort.isFunction())
37    {
38      sort = sort.getFunctionCodomainSort();
39    }
40    StringBuilder ss = new StringBuilder();
41    ss.append("(define-fun ").append(f).append(" (");
42    for (int i = 0; i < params.length; ++i)
43    {
44      if (i > 0)
45      {
46        ss.append(' ');
47      }
48      ss.append('(').append(params[i]).append(' ').append(params[i].getSort()).append(')');
49    }
50    ss.append(") ").append(sort).append(' ').append(body).append(')');
51    return ss.toString();
52  }
53
54  /**
55   * Print solutions for synthesis conjecture to the standard output stream.
56   *
57   * @param terms the terms for which the synthesis solutions were retrieved
58   * @param sols the synthesis solutions of the given terms
59   */
60  public static void printSynthSolutions(Term[] terms, Term[] sols) throws CVC5ApiException
61  {
62    System.out.println('(');
63    for (int i = 0; i < terms.length; ++i)
64    {
65      List<Term> params = new ArrayList<>();
66      Term body = sols[i];
67      if (sols[i].getKind() == Kind.LAMBDA)
68      {
69        for (Term t : sols[i].getChild(0))
70        {
71          params.add(t);
72        }
73        body = sols[i].getChild(1);
74      }
75      System.out.println("  " + defineFunToString(terms[i], params.toArray(new Term[0]), body));
76    }
77    System.out.println(')');
78  }
79}
 1#!/usr/bin/env python
 2###############################################################################
 3# Top contributors (to current version):
 4#   Yoni Zohar, Abdalrhman Mohamed, Alex Ozdemir
 5#
 6# This file is part of the cvc5 project.
 7#
 8# Copyright (c) 2009-2025 by the authors listed in the file AUTHORS
 9# in the top-level source directory and their institutional affiliations.
10# All rights reserved.  See the file COPYING in the top-level source
11# directory for licensing information.
12# #############################################################################
13#
14# Utility Methods, translated from examples/api/utils.h
15##
16
17import cvc5
18from cvc5 import Kind
19
20# Get the string version of define-fun command.
21# @param f the function to print
22# @param params the function parameters
23# @param body the function body
24# @return a string version of define-fun
25
26
27def define_fun_to_string(f, params, body):
28    sort = f.getSort()
29    if sort.isFunction():
30        sort = f.getSort().getFunctionCodomainSort()
31    result = "(define-fun " + str(f) + " ("
32    for i in range(0, len(params)):
33        if i > 0:
34            result += " "
35        result += "(" + str(params[i]) + " " + str(params[i].getSort()) + ")"
36    result += ") " + str(sort) + " " + str(body) + ")"
37    return result
38
39
40# Print solutions for synthesis conjecture to the standard output stream.
41# @param terms the terms for which the synthesis solutions were retrieved
42# @param sols the synthesis solutions of the given terms
43
44
45def print_synth_solutions(terms, sols):
46    result = "(\n"
47    for i in range(0, len(terms)):
48        params = []
49        body = sols[i]
50        if sols[i].getKind() == Kind.LAMBDA:
51            params += sols[i][0]
52            body = sols[i][1]
53        result += "  " + define_fun_to_string(terms[i], params, body) + "\n"
54    result += ")"
55    print(result)