CBMC
smt2_conv.cpp
Go to the documentation of this file.
1 /*******************************************************************\
2 
3 Module: SMT Backend
4 
5 Author: Daniel Kroening, kroening@kroening.com
6 
7 \*******************************************************************/
8 
11 
12 #include "smt2_conv.h"
13 
14 #include <util/arith_tools.h>
15 #include <util/bitvector_expr.h>
16 #include <util/byte_operators.h>
17 #include <util/c_types.h>
18 #include <util/config.h>
19 #include <util/expr_iterator.h>
20 #include <util/expr_util.h>
21 #include <util/fixedbv.h>
22 #include <util/floatbv_expr.h>
23 #include <util/format_expr.h>
24 #include <util/ieee_float.h>
25 #include <util/invariant.h>
26 #include <util/mathematical_expr.h>
27 #include <util/namespace.h>
30 #include <util/prefix.h>
31 #include <util/range.h>
32 #include <util/simplify_expr.h>
33 #include <util/std_expr.h>
34 #include <util/string2int.h>
35 #include <util/string_constant.h>
36 #include <util/threeval.h>
37 
43 
44 #include "smt2_tokenizer.h"
45 
46 // Mark different kinds of error conditions
47 
48 // Unexpected types and other combinations not implemented and not
49 // expected to be needed
50 #define UNEXPECTEDCASE(S) PRECONDITION_WITH_DIAGNOSTICS(false, S);
51 
52 // General todos
53 #define SMT2_TODO(S) PRECONDITION_WITH_DIAGNOSTICS(false, "TODO: " S)
54 
56  const namespacet &_ns,
57  const std::string &_benchmark,
58  const std::string &_notes,
59  const std::string &_logic,
60  solvert _solver,
61  std::ostream &_out)
62  : use_FPA_theory(false),
63  use_array_of_bool(false),
64  use_as_const(false),
65  use_check_sat_assuming(false),
66  use_datatypes(false),
67  use_lambda_for_array(false),
68  emit_set_logic(true),
69  ns(_ns),
70  out(_out),
71  benchmark(_benchmark),
72  notes(_notes),
73  logic(_logic),
74  solver(_solver),
75  boolbv_width(_ns),
76  pointer_logic(_ns),
77  no_boolean_variables(0)
78 {
79  // We set some defaults differently
80  // for some solvers.
81 
82  switch(solver)
83  {
84  case solvert::GENERIC:
85  break;
86 
87  case solvert::BOOLECTOR:
88  break;
89 
91  use_FPA_theory = true;
92  use_array_of_bool = true;
93  use_as_const = true;
95  emit_set_logic = false;
96  break;
97 
98  case solvert::CVC3:
99  break;
100 
101  case solvert::CVC4:
102  logic = "ALL";
103  use_array_of_bool = true;
104  use_as_const = true;
105  break;
106 
107  case solvert::MATHSAT:
108  break;
109 
110  case solvert::YICES:
111  break;
112 
113  case solvert::Z3:
114  use_array_of_bool = true;
115  use_as_const = true;
116  use_check_sat_assuming = true;
117  use_lambda_for_array = true;
118  emit_set_logic = false;
119  use_datatypes = true;
120  break;
121  }
122 
123  write_header();
124 }
125 
127 {
128  return "SMT2";
129 }
130 
131 void smt2_convt::print_assignment(std::ostream &os) const
132 {
133  // Boolean stuff
134 
135  for(std::size_t v=0; v<boolean_assignment.size(); v++)
136  os << "b" << v << "=" << boolean_assignment[v] << "\n";
137 
138  // others
139 }
140 
142 {
143  if(l.is_true())
144  return tvt(true);
145  if(l.is_false())
146  return tvt(false);
147 
148  INVARIANT(
149  l.var_no() < boolean_assignment.size(),
150  "variable number shall be within bounds");
151  return tvt(boolean_assignment[l.var_no()]^l.sign());
152 }
153 
155 {
156  out << "; SMT 2" << "\n";
157 
158  switch(solver)
159  {
160  // clang-format off
161  case solvert::GENERIC: break;
162  case solvert::BOOLECTOR: out << "; Generated for Boolector\n"; break;
164  out << "; Generated for the CPROVER SMT2 solver\n"; break;
165  case solvert::CVC3: out << "; Generated for CVC 3\n"; break;
166  case solvert::CVC4: out << "; Generated for CVC 4\n"; break;
167  case solvert::MATHSAT: out << "; Generated for MathSAT\n"; break;
168  case solvert::YICES: out << "; Generated for Yices\n"; break;
169  case solvert::Z3: out << "; Generated for Z3\n"; break;
170  // clang-format on
171  }
172 
173  out << "(set-info :source \"" << notes << "\")" << "\n";
174 
175  out << "(set-option :produce-models true)" << "\n";
176 
177  // We use a broad mixture of logics, so on some solvers
178  // its better not to declare here.
179  // set-logic should come after setting options
180  if(emit_set_logic && !logic.empty())
181  out << "(set-logic " << logic << ")" << "\n";
182 }
183 
185 {
186  out << "\n";
187 
188  // fix up the object sizes
189  for(const auto &object : object_sizes)
190  define_object_size(object.second, object.first);
191 
192  if(use_check_sat_assuming && !assumptions.empty())
193  {
194  out << "(check-sat-assuming (";
195  for(const auto &assumption : assumptions)
196  convert_literal(to_literal_expr(assumption).get_literal());
197  out << "))\n";
198  }
199  else
200  {
201  // add the assumptions, if any
202  if(!assumptions.empty())
203  {
204  out << "; assumptions\n";
205 
206  for(const auto &assumption : assumptions)
207  {
208  out << "(assert ";
209  convert_literal(to_literal_expr(assumption).get_literal());
210  out << ")"
211  << "\n";
212  }
213  }
214 
215  out << "(check-sat)\n";
216  }
217 
218  out << "\n";
219 
221  {
222  for(const auto &id : smt2_identifiers)
223  out << "(get-value (" << id << "))"
224  << "\n";
225  }
226 
227  out << "\n";
228 
229  out << "(exit)\n";
230 
231  out << "; end of SMT2 file"
232  << "\n";
233 }
234 
236  const irep_idt &id,
237  const object_size_exprt &expr)
238 {
239  const exprt &ptr = expr.pointer();
240  std::size_t size_width = boolbv_width(expr.type());
241  std::size_t pointer_width = boolbv_width(ptr.type());
242  std::size_t number = 0;
243  std::size_t h=pointer_width-1;
244  std::size_t l=pointer_width-config.bv_encoding.object_bits;
245 
246  for(const auto &o : pointer_logic.objects)
247  {
248  const typet &type = o.type();
249  auto size_expr = size_of_expr(type, ns);
250  const auto object_size =
251  numeric_cast<mp_integer>(size_expr.value_or(nil_exprt()));
252 
253  if(
254  (o.id() != ID_symbol && o.id() != ID_string_constant) ||
255  !size_expr.has_value() || !object_size.has_value())
256  {
257  ++number;
258  continue;
259  }
260 
261  out << "(assert (=> (= "
262  << "((_ extract " << h << " " << l << ") ";
263  convert_expr(ptr);
264  out << ") (_ bv" << number << " " << config.bv_encoding.object_bits << "))"
265  << "(= " << id << " (_ bv" << *object_size << " " << size_width
266  << "))))\n";
267 
268  ++number;
269  }
270 }
271 
273 {
274  write_footer();
275  out.flush();
277 }
278 
279 exprt smt2_convt::get(const exprt &expr) const
280 {
281  if(expr.id()==ID_symbol)
282  {
283  const irep_idt &id=to_symbol_expr(expr).get_identifier();
284 
285  identifier_mapt::const_iterator it=identifier_map.find(id);
286 
287  if(it!=identifier_map.end())
288  return it->second.value;
289  return expr;
290  }
291  else if(expr.id()==ID_nondet_symbol)
292  {
294 
295  identifier_mapt::const_iterator it=identifier_map.find(id);
296 
297  if(it!=identifier_map.end())
298  return it->second.value;
299  }
300  else if(expr.id()==ID_member)
301  {
302  const member_exprt &member_expr=to_member_expr(expr);
303  exprt tmp=get(member_expr.struct_op());
304  if(tmp.is_nil())
305  return nil_exprt();
306  return member_exprt(tmp, member_expr.get_component_name(), expr.type());
307  }
308  else if(expr.id() == ID_literal)
309  {
310  auto l = to_literal_expr(expr).get_literal();
311  if(l_get(l).is_true())
312  return true_exprt();
313  else
314  return false_exprt();
315  }
316  else if(expr.id() == ID_not)
317  {
318  auto op = get(to_not_expr(expr).op());
319  if(op.is_true())
320  return false_exprt();
321  else if(op.is_false())
322  return true_exprt();
323  }
324  else if(expr.is_constant())
325  return expr;
326  else if(const auto &array = expr_try_dynamic_cast<array_exprt>(expr))
327  {
328  exprt array_copy = *array;
329  for(auto &element : array_copy.operands())
330  {
331  element = get(element);
332  }
333  return array_copy;
334  }
335 
336  return nil_exprt();
337 }
338 
340  const irept &src,
341  const typet &type)
342 {
343  // See http://www.grammatech.com/resources/smt/SMTLIBTutorial.pdf for the
344  // syntax of SMTlib2 literals.
345  //
346  // A literal expression is one of the following forms:
347  //
348  // * Numeral -- this is a natural number in decimal and is of the form:
349  // 0|([1-9][0-9]*)
350  // * Decimal -- this is a decimal expansion of a real number of the form:
351  // (0|[1-9][0-9]*)[.]([0-9]+)
352  // * Binary -- this is a natural number in binary and is of the form:
353  // #b[01]+
354  // * Hex -- this is a natural number in hexadecimal and is of the form:
355  // #x[0-9a-fA-F]+
356  //
357  // Right now I'm not parsing decimals. It'd be nice if we had a real YACC
358  // parser here, but whatever.
359 
360  mp_integer value;
361 
362  if(!src.id().empty())
363  {
364  const std::string &s=src.id_string();
365 
366  if(s.size()>=2 && s[0]=='#' && s[1]=='b')
367  {
368  // Binary #b010101
369  value=string2integer(s.substr(2), 2);
370  }
371  else if(s.size()>=2 && s[0]=='#' && s[1]=='x')
372  {
373  // Hex #x012345
374  value=string2integer(s.substr(2), 16);
375  }
376  else
377  {
378  // Numeral
379  value=string2integer(s);
380  }
381  }
382  else if(src.get_sub().size()==2 &&
383  src.get_sub()[0].id()=="-") // (- 100)
384  {
385  value=-string2integer(src.get_sub()[1].id_string());
386  }
387  else if(src.get_sub().size()==3 &&
388  src.get_sub()[0].id()=="_" &&
389  // (_ bvDECIMAL_VALUE SIZE)
390  src.get_sub()[1].id_string().substr(0, 2)=="bv")
391  {
392  value=string2integer(src.get_sub()[1].id_string().substr(2));
393  }
394  else if(src.get_sub().size()==4 &&
395  src.get_sub()[0].id()=="fp") // (fp signbv exponentbv significandbv)
396  {
397  if(type.id()==ID_floatbv)
398  {
399  const floatbv_typet &floatbv_type=to_floatbv_type(type);
402  parse_literal(src.get_sub()[2], unsignedbv_typet(floatbv_type.get_e()));
403  constant_exprt s3 =
404  parse_literal(src.get_sub()[3], unsignedbv_typet(floatbv_type.get_f()));
405 
406  const auto s1_int = numeric_cast_v<mp_integer>(s1);
407  const auto s2_int = numeric_cast_v<mp_integer>(s2);
408  const auto s3_int = numeric_cast_v<mp_integer>(s3);
409 
410  // stitch the bits together
411  value = bitwise_or(
412  s1_int << (floatbv_type.get_e() + floatbv_type.get_f()),
413  bitwise_or((s2_int << floatbv_type.get_f()), s3_int));
414  }
415  else
416  value=0;
417  }
418  else if(src.get_sub().size()==4 &&
419  src.get_sub()[0].id()=="_" &&
420  src.get_sub()[1].id()=="+oo") // (_ +oo e s)
421  {
422  std::size_t e = unsafe_string2size_t(src.get_sub()[2].id_string());
423  std::size_t s = unsafe_string2size_t(src.get_sub()[3].id_string());
425  }
426  else if(src.get_sub().size()==4 &&
427  src.get_sub()[0].id()=="_" &&
428  src.get_sub()[1].id()=="-oo") // (_ -oo e s)
429  {
430  std::size_t e = unsafe_string2size_t(src.get_sub()[2].id_string());
431  std::size_t s = unsafe_string2size_t(src.get_sub()[3].id_string());
433  }
434  else if(src.get_sub().size()==4 &&
435  src.get_sub()[0].id()=="_" &&
436  src.get_sub()[1].id()=="NaN") // (_ NaN e s)
437  {
438  std::size_t e = unsafe_string2size_t(src.get_sub()[2].id_string());
439  std::size_t s = unsafe_string2size_t(src.get_sub()[3].id_string());
440  return ieee_floatt::NaN(ieee_float_spect(s - 1, e)).to_expr();
441  }
442 
443  if(type.id()==ID_signedbv ||
444  type.id()==ID_unsignedbv ||
445  type.id()==ID_c_enum ||
446  type.id()==ID_c_bool)
447  {
448  return from_integer(value, type);
449  }
450  else if(type.id()==ID_c_enum_tag)
451  {
452  constant_exprt result =
454 
455  // restore the c_enum_tag type
456  result.type() = type;
457  return result;
458  }
459  else if(type.id()==ID_fixedbv ||
460  type.id()==ID_floatbv)
461  {
462  std::size_t width=boolbv_width(type);
463  return constant_exprt(integer2bvrep(value, width), type);
464  }
465  else if(type.id()==ID_integer ||
466  type.id()==ID_range)
467  {
468  return from_integer(value, type);
469  }
470  else
471  INVARIANT(
472  false,
473  "smt2_convt::parse_literal should not be of unsupported type " +
474  type.id_string());
475 }
476 
478  const irept &src,
479  const array_typet &type)
480 {
481  std::unordered_map<int64_t, exprt> operands_map;
482  walk_array_tree(&operands_map, src, type);
483  exprt::operandst operands;
484  // Try to find the default value, if there is none then set it
485  auto maybe_default_op = operands_map.find(-1);
486  exprt default_op;
487  if(maybe_default_op == operands_map.end())
488  default_op = nil_exprt();
489  else
490  default_op = maybe_default_op->second;
491  int64_t i = 0;
492  auto maybe_size = numeric_cast<std::int64_t>(type.size());
493  if(maybe_size.has_value())
494  {
495  while(i < maybe_size.value())
496  {
497  auto found_op = operands_map.find(i);
498  if(found_op != operands_map.end())
499  operands.emplace_back(found_op->second);
500  else
501  operands.emplace_back(default_op);
502  i++;
503  }
504  }
505  else
506  {
507  // Array size is unknown, keep adding with known indexes in order
508  // until we fail to find one.
509  auto found_op = operands_map.find(i);
510  while(found_op != operands_map.end())
511  {
512  operands.emplace_back(found_op->second);
513  i++;
514  found_op = operands_map.find(i);
515  }
516  operands.emplace_back(default_op);
517  }
518  return array_exprt(operands, type);
519 }
520 
522  std::unordered_map<int64_t, exprt> *operands_map,
523  const irept &src,
524  const array_typet &type)
525 {
526  if(src.get_sub().size()==4 && src.get_sub()[0].id()=="store")
527  {
528  // This is the SMT syntax being parsed here
529  // (store array index value)
530  // Recurse
531  walk_array_tree(operands_map, src.get_sub()[1], type);
532  const auto index_expr = parse_rec(src.get_sub()[2], type.size().type());
533  const constant_exprt index_constant = to_constant_expr(index_expr);
534  mp_integer tempint;
535  bool failure = to_integer(index_constant, tempint);
536  if(failure)
537  return;
538  long index = tempint.to_long();
539  exprt value = parse_rec(src.get_sub()[3], type.element_type());
540  operands_map->emplace(index, value);
541  }
542  else if(src.get_sub().size() == 3 && src.get_sub()[0].id() == "let")
543  {
544  // This is produced by Z3
545  // (let (....) (....))
547  operands_map, src.get_sub()[1].get_sub()[0].get_sub()[1], type);
548  walk_array_tree(operands_map, src.get_sub()[2], type);
549  }
550  else if(src.get_sub().size()==2 &&
551  src.get_sub()[0].get_sub().size()==3 &&
552  src.get_sub()[0].get_sub()[0].id()=="as" &&
553  src.get_sub()[0].get_sub()[1].id()=="const")
554  {
555  // (as const type_info default_value)
556  exprt default_value = parse_rec(src.get_sub()[1], type.element_type());
557  operands_map->emplace(-1, default_value);
558  }
559 }
560 
562  const irept &src,
563  const union_typet &type)
564 {
565  // these are always flat
566  PRECONDITION(!type.components().empty());
567  const union_typet::componentt &first=type.components().front();
568  std::size_t width=boolbv_width(type);
569  exprt value = parse_rec(src, unsignedbv_typet(width));
570  if(value.is_nil())
571  return nil_exprt();
572  const typecast_exprt converted(value, first.type());
573  return union_exprt(first.get_name(), converted, type);
574 }
575 
578 {
579  const struct_typet::componentst &components =
580  type.components();
581 
582  struct_exprt result(exprt::operandst(components.size(), nil_exprt()), type);
583 
584  if(components.empty())
585  return result;
586 
587  if(use_datatypes)
588  {
589  // Structs look like:
590  // (mk-struct.1 <component0> <component1> ... <componentN>)
591 
592  if(src.get_sub().size()!=components.size()+1)
593  return result; // give up
594 
595  for(std::size_t i=0; i<components.size(); i++)
596  {
597  const struct_typet::componentt &c=components[i];
598  result.operands()[i]=parse_rec(src.get_sub()[i+1], c.type());
599  }
600  }
601  else
602  {
603  // These are just flattened, i.e., we expect to see a monster bit vector.
604  std::size_t total_width=boolbv_width(type);
605  const auto l = parse_literal(src, unsignedbv_typet(total_width));
606 
607  const irep_idt binary =
608  integer2binary(numeric_cast_v<mp_integer>(l), total_width);
609 
610  CHECK_RETURN(binary.size() == total_width);
611 
612  std::size_t offset=0;
613 
614  for(std::size_t i=0; i<components.size(); i++)
615  {
616  std::size_t component_width=boolbv_width(components[i].type());
617 
618  INVARIANT(
619  offset + component_width <= total_width,
620  "struct component bits shall be within struct bit vector");
621 
622  std::string component_binary=
623  "#b"+id2string(binary).substr(
624  total_width-offset-component_width, component_width);
625 
626  result.operands()[i]=
627  parse_rec(irept(component_binary), components[i].type());
628 
629  offset+=component_width;
630  }
631  }
632 
633  return result;
634 }
635 
636 exprt smt2_convt::parse_rec(const irept &src, const typet &type)
637 {
638  if(
639  type.id() == ID_signedbv || type.id() == ID_unsignedbv ||
640  type.id() == ID_integer || type.id() == ID_rational ||
641  type.id() == ID_real || type.id() == ID_c_enum ||
642  type.id() == ID_c_enum_tag || type.id() == ID_fixedbv ||
643  type.id() == ID_floatbv)
644  {
645  return parse_literal(src, type);
646  }
647  else if(type.id()==ID_bool)
648  {
649  if(src.id()==ID_1 || src.id()==ID_true)
650  return true_exprt();
651  else if(src.id()==ID_0 || src.id()==ID_false)
652  return false_exprt();
653  }
654  else if(type.id()==ID_pointer)
655  {
656  // these come in as bit-vector literals
657  std::size_t width=boolbv_width(type);
658  constant_exprt bv_expr = parse_literal(src, unsignedbv_typet(width));
659 
660  mp_integer v = numeric_cast_v<mp_integer>(bv_expr);
661 
662  // split into object and offset
665  ptr.object = numeric_cast_v<std::size_t>(v / pow);
666  ptr.offset=v%pow;
668  bv_expr.get_value(),
670  }
671  else if(type.id()==ID_struct)
672  {
673  return parse_struct(src, to_struct_type(type));
674  }
675  else if(type.id() == ID_struct_tag)
676  {
677  auto struct_expr =
679  // restore the tag type
680  struct_expr.type() = type;
681  return std::move(struct_expr);
682  }
683  else if(type.id()==ID_union)
684  {
685  return parse_union(src, to_union_type(type));
686  }
687  else if(type.id() == ID_union_tag)
688  {
689  auto union_expr = parse_union(src, ns.follow_tag(to_union_tag_type(type)));
690  // restore the tag type
691  union_expr.type() = type;
692  return union_expr;
693  }
694  else if(type.id()==ID_array)
695  {
696  return parse_array(src, to_array_type(type));
697  }
698 
699  return nil_exprt();
700 }
701 
703  const exprt &expr,
704  const pointer_typet &result_type)
705 {
706  if(expr.id()==ID_symbol ||
707  expr.id()==ID_constant ||
708  expr.id()==ID_string_constant ||
709  expr.id()==ID_label)
710  {
711  out
712  << "(concat (_ bv"
713  << pointer_logic.add_object(expr) << " "
714  << config.bv_encoding.object_bits << ")"
715  << " (_ bv0 "
716  << boolbv_width(result_type)-config.bv_encoding.object_bits << "))";
717  }
718  else if(expr.id()==ID_index)
719  {
720  const index_exprt &index_expr = to_index_expr(expr);
721 
722  const exprt &array = index_expr.array();
723  const exprt &index = index_expr.index();
724 
725  if(index.is_zero())
726  {
727  if(array.type().id()==ID_pointer)
728  convert_expr(array);
729  else if(array.type().id()==ID_array)
730  convert_address_of_rec(array, result_type);
731  else
732  UNREACHABLE;
733  }
734  else
735  {
736  // this is really pointer arithmetic
737  index_exprt new_index_expr = index_expr;
738  new_index_expr.index() = from_integer(0, index.type());
739 
740  address_of_exprt address_of_expr(
741  new_index_expr,
743 
744  plus_exprt plus_expr{address_of_expr, index};
745 
746  convert_expr(plus_expr);
747  }
748  }
749  else if(expr.id()==ID_member)
750  {
751  const member_exprt &member_expr=to_member_expr(expr);
752 
753  const exprt &struct_op=member_expr.struct_op();
754  const typet &struct_op_type = struct_op.type();
755 
757  struct_op_type.id() == ID_struct || struct_op_type.id() == ID_struct_tag,
758  "member expression operand shall have struct type");
759 
760  const struct_typet &struct_type = to_struct_type(ns.follow(struct_op_type));
761 
762  const irep_idt &component_name = member_expr.get_component_name();
763 
764  const auto offset = member_offset(struct_type, component_name, ns);
765  CHECK_RETURN(offset.has_value() && *offset >= 0);
766 
768 
769  // pointer arithmetic
770  out << "(bvadd ";
771  convert_address_of_rec(struct_op, result_type);
773  out << ")"; // bvadd
774  }
775  else if(expr.id()==ID_if)
776  {
777  const if_exprt &if_expr = to_if_expr(expr);
778 
779  out << "(ite ";
780  convert_expr(if_expr.cond());
781  out << " ";
782  convert_address_of_rec(if_expr.true_case(), result_type);
783  out << " ";
784  convert_address_of_rec(if_expr.false_case(), result_type);
785  out << ")";
786  }
787  else
788  INVARIANT(
789  false,
790  "operand of address of expression should not be of kind " +
791  expr.id_string());
792 }
793 
794 static bool has_quantifier(const exprt &expr)
795 {
796  bool result = false;
797  expr.visit_post([&result](const exprt &node) {
798  if(node.id() == ID_exists || node.id() == ID_forall)
799  result = true;
800  });
801  return result;
802 }
803 
805 {
806  PRECONDITION(expr.type().id() == ID_bool);
807 
808  // Three cases where no new handle is needed.
809 
810  if(expr.is_true())
811  return const_literal(true);
812  else if(expr.is_false())
813  return const_literal(false);
814  else if(expr.id()==ID_literal)
815  return to_literal_expr(expr).get_literal();
816 
817  // Need a new handle
818 
819  out << "\n";
820 
821  exprt prepared_expr = prepare_for_convert_expr(expr);
822 
823  literalt l(no_boolean_variables, false);
825 
826  out << "; convert\n";
827  out << "; Converting var_no " << l.var_no() << " with expr ID of "
828  << expr.id_string() << "\n";
829  // We're converting the expression, so store it in the defined_expressions
830  // store and in future we use the literal instead of the whole expression
831  // Note that here we are always converting, so we do not need to consider
832  // other literal kinds, only "|B###|"
833 
834  // Z3 refuses get-value when a defined symbol contains a quantifier.
835  if(has_quantifier(prepared_expr))
836  {
837  out << "(declare-fun ";
838  convert_literal(l);
839  out << " () Bool)\n";
840  out << "(assert (= ";
841  convert_literal(l);
842  out << ' ';
843  convert_expr(prepared_expr);
844  out << "))\n";
845  }
846  else
847  {
848  auto identifier =
849  convert_identifier(std::string{"B"} + std::to_string(l.var_no()));
850  defined_expressions[expr] = identifier;
851  smt2_identifiers.insert(identifier);
852  out << "(define-fun " << identifier << " () Bool ";
853  convert_expr(prepared_expr);
854  out << ")\n";
855  }
856 
857  return l;
858 }
859 
861 {
862  // We can only improve Booleans.
863  if(expr.type().id() != ID_bool)
864  return expr;
865 
866  return literal_exprt(convert(expr));
867 }
868 
870 {
871  if(l==const_literal(false))
872  out << "false";
873  else if(l==const_literal(true))
874  out << "true";
875  else
876  {
877  if(l.sign())
878  out << "(not ";
879 
880  const auto identifier =
882 
883  out << identifier;
884 
885  if(l.sign())
886  out << ")";
887 
888  smt2_identifiers.insert(identifier);
889  }
890 }
891 
893 {
895 }
896 
897 void smt2_convt::push(const std::vector<exprt> &_assumptions)
898 {
899  INVARIANT(assumptions.empty(), "nested contexts are not supported");
900 
901  assumptions = _assumptions;
902 }
903 
905 {
906  assumptions.clear();
907 }
908 
909 static bool is_smt2_simple_identifier(const std::string &identifier)
910 {
911  if(identifier.empty())
912  return false;
913 
914  if(isdigit(identifier[0]))
915  return false;
916 
917  for(auto ch : id2string(identifier))
918  {
920  return false;
921  }
922 
923  return true;
924 }
925 
926 std::string smt2_convt::convert_identifier(const irep_idt &identifier)
927 {
928  // Is this a "simple identifier"?
929  if(is_smt2_simple_identifier(id2string(identifier)))
930  return id2string(identifier);
931 
932  // Backslashes are disallowed in quoted symbols just for simplicity.
933  // Otherwise, for Common Lisp compatibility they would have to be treated
934  // as escaping symbols.
935 
936  std::string result = "|";
937 
938  for(auto ch : identifier)
939  {
940  switch(ch)
941  {
942  case '|':
943  case '\\':
944  case '&': // we use the & for escaping
945  result+='&';
946  result+=std::to_string(ch);
947  result+=';';
948  break;
949 
950  case '$': // $ _is_ allowed
951  default:
952  result+=ch;
953  }
954  }
955 
956  result += '|';
957 
958  return result;
959 }
960 
961 std::string smt2_convt::type2id(const typet &type) const
962 {
963  if(type.id()==ID_floatbv)
964  {
965  ieee_float_spect spec(to_floatbv_type(type));
966  return "f"+std::to_string(spec.width())+"_"+std::to_string(spec.f);
967  }
968  else if(type.id()==ID_unsignedbv)
969  {
970  return "u"+std::to_string(to_unsignedbv_type(type).get_width());
971  }
972  else if(type.id()==ID_c_bool)
973  {
974  return "u"+std::to_string(to_c_bool_type(type).get_width());
975  }
976  else if(type.id()==ID_signedbv)
977  {
978  return "s"+std::to_string(to_signedbv_type(type).get_width());
979  }
980  else if(type.id()==ID_bool)
981  {
982  return "b";
983  }
984  else if(type.id()==ID_c_enum_tag)
985  {
986  return type2id(ns.follow_tag(to_c_enum_tag_type(type)).underlying_type());
987  }
988  else if(type.id() == ID_pointer)
989  {
990  return "p" + std::to_string(to_pointer_type(type).get_width());
991  }
992  else if(type.id() == ID_struct_tag)
993  {
994  if(use_datatypes)
995  return datatype_map.at(type);
996  else
997  return "S" + std::to_string(boolbv_width(type));
998  }
999  else
1000  {
1001  UNREACHABLE;
1002  }
1003 }
1004 
1005 std::string smt2_convt::floatbv_suffix(const exprt &expr) const
1006 {
1007  PRECONDITION(!expr.operands().empty());
1008  return "_" + type2id(to_multi_ary_expr(expr).op0().type()) + "->" +
1009  type2id(expr.type());
1010 }
1011 
1013 {
1015 
1016  if(expr.id()==ID_symbol)
1017  {
1018  const irep_idt &id = to_symbol_expr(expr).get_identifier();
1019  out << convert_identifier(id);
1020  return;
1021  }
1022 
1023  if(expr.id()==ID_smt2_symbol)
1024  {
1025  const irep_idt &id = to_smt2_symbol(expr).get_identifier();
1026  out << id;
1027  return;
1028  }
1029 
1030  INVARIANT(
1031  !expr.operands().empty(), "non-symbol expressions shall have operands");
1032 
1033  out << '('
1034  << convert_identifier(
1035  "float_bv." + expr.id_string() + floatbv_suffix(expr));
1036 
1037  forall_operands(it, expr)
1038  {
1039  out << ' ';
1040  convert_expr(*it);
1041  }
1042 
1043  out << ')';
1044 }
1045 
1047 {
1048  // huge monster case split over expression id
1049  if(expr.id()==ID_symbol)
1050  {
1051  const irep_idt &id = to_symbol_expr(expr).get_identifier();
1052  DATA_INVARIANT(!id.empty(), "symbol must have identifier");
1053  out << convert_identifier(id);
1054  }
1055  else if(expr.id()==ID_nondet_symbol)
1056  {
1057  const irep_idt &id = to_nondet_symbol_expr(expr).get_identifier();
1058  DATA_INVARIANT(!id.empty(), "nondet symbol must have identifier");
1059  out << convert_identifier("nondet_" + id2string(id));
1060  }
1061  else if(expr.id()==ID_smt2_symbol)
1062  {
1063  const irep_idt &id = to_smt2_symbol(expr).get_identifier();
1064  DATA_INVARIANT(!id.empty(), "smt2 symbol must have identifier");
1065  out << id;
1066  }
1067  else if(expr.id()==ID_typecast)
1068  {
1070  }
1071  else if(expr.id()==ID_floatbv_typecast)
1072  {
1074  }
1075  else if(expr.id()==ID_struct)
1076  {
1078  }
1079  else if(expr.id()==ID_union)
1080  {
1082  }
1083  else if(expr.id()==ID_constant)
1084  {
1086  }
1087  else if(expr.id() == ID_concatenation && expr.operands().size() == 1)
1088  {
1090  expr.type() == expr.operands().front().type(),
1091  "concatenation over a single operand should have matching types",
1092  expr.pretty());
1093 
1094  convert_expr(expr.operands().front());
1095  }
1096  else if(expr.id()==ID_concatenation ||
1097  expr.id()==ID_bitand ||
1098  expr.id()==ID_bitor ||
1099  expr.id()==ID_bitxor ||
1100  expr.id()==ID_bitnand ||
1101  expr.id()==ID_bitnor)
1102  {
1104  expr.operands().size() >= 2,
1105  "given expression should have at least two operands",
1106  expr.id_string());
1107 
1108  out << "(";
1109 
1110  if(expr.id()==ID_concatenation)
1111  out << "concat";
1112  else if(expr.id()==ID_bitand)
1113  out << "bvand";
1114  else if(expr.id()==ID_bitor)
1115  out << "bvor";
1116  else if(expr.id()==ID_bitxor)
1117  out << "bvxor";
1118  else if(expr.id()==ID_bitnand)
1119  out << "bvnand";
1120  else if(expr.id()==ID_bitnor)
1121  out << "bvnor";
1122 
1123  forall_operands(it, expr)
1124  {
1125  out << " ";
1126  flatten2bv(*it);
1127  }
1128 
1129  out << ")";
1130  }
1131  else if(expr.id()==ID_bitnot)
1132  {
1133  const bitnot_exprt &bitnot_expr = to_bitnot_expr(expr);
1134 
1135  if(bitnot_expr.type().id() == ID_vector)
1136  {
1137  if(use_datatypes)
1138  {
1139  const std::string &smt_typename = datatype_map.at(bitnot_expr.type());
1140 
1141  // extract elements
1142  const vector_typet &vector_type = to_vector_type(bitnot_expr.type());
1143 
1144  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
1145 
1146  out << "(let ((?vectorop ";
1147  convert_expr(bitnot_expr.op());
1148  out << ")) ";
1149 
1150  out << "(mk-" << smt_typename;
1151 
1152  typet index_type=vector_type.size().type();
1153 
1154  // do bitnot component-by-component
1155  for(mp_integer i=0; i!=size; ++i)
1156  {
1157  out << " (bvnot (" << smt_typename << "." << (size-i-1)
1158  << " ?vectorop))";
1159  }
1160 
1161  out << "))"; // mk-, let
1162  }
1163  else
1164  SMT2_TODO("bitnot for vectors");
1165  }
1166  else
1167  {
1168  out << "(bvnot ";
1169  convert_expr(bitnot_expr.op());
1170  out << ")";
1171  }
1172  }
1173  else if(expr.id()==ID_unary_minus)
1174  {
1175  const unary_minus_exprt &unary_minus_expr = to_unary_minus_expr(expr);
1176 
1177  if(
1178  unary_minus_expr.type().id() == ID_rational ||
1179  unary_minus_expr.type().id() == ID_integer ||
1180  unary_minus_expr.type().id() == ID_real)
1181  {
1182  out << "(- ";
1183  convert_expr(unary_minus_expr.op());
1184  out << ")";
1185  }
1186  else if(unary_minus_expr.type().id() == ID_floatbv)
1187  {
1188  // this has no rounding mode
1189  if(use_FPA_theory)
1190  {
1191  out << "(fp.neg ";
1192  convert_expr(unary_minus_expr.op());
1193  out << ")";
1194  }
1195  else
1196  convert_floatbv(unary_minus_expr);
1197  }
1198  else if(unary_minus_expr.type().id() == ID_vector)
1199  {
1200  if(use_datatypes)
1201  {
1202  const std::string &smt_typename =
1203  datatype_map.at(unary_minus_expr.type());
1204 
1205  // extract elements
1206  const vector_typet &vector_type =
1207  to_vector_type(unary_minus_expr.type());
1208 
1209  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
1210 
1211  out << "(let ((?vectorop ";
1212  convert_expr(unary_minus_expr.op());
1213  out << ")) ";
1214 
1215  out << "(mk-" << smt_typename;
1216 
1217  typet index_type=vector_type.size().type();
1218 
1219  // negate component-by-component
1220  for(mp_integer i=0; i!=size; ++i)
1221  {
1222  out << " (bvneg (" << smt_typename << "." << (size-i-1)
1223  << " ?vectorop))";
1224  }
1225 
1226  out << "))"; // mk-, let
1227  }
1228  else
1229  SMT2_TODO("unary minus for vector");
1230  }
1231  else
1232  {
1233  out << "(bvneg ";
1234  convert_expr(unary_minus_expr.op());
1235  out << ")";
1236  }
1237  }
1238  else if(expr.id()==ID_unary_plus)
1239  {
1240  // A no-op (apart from type promotion)
1241  convert_expr(to_unary_plus_expr(expr).op());
1242  }
1243  else if(expr.id()==ID_sign)
1244  {
1245  const sign_exprt &sign_expr = to_sign_expr(expr);
1246 
1247  const typet &op_type = sign_expr.op().type();
1248 
1249  if(op_type.id()==ID_floatbv)
1250  {
1251  if(use_FPA_theory)
1252  {
1253  out << "(fp.isNegative ";
1254  convert_expr(sign_expr.op());
1255  out << ")";
1256  }
1257  else
1258  convert_floatbv(sign_expr);
1259  }
1260  else if(op_type.id()==ID_signedbv)
1261  {
1262  std::size_t op_width=to_signedbv_type(op_type).get_width();
1263 
1264  out << "(bvslt ";
1265  convert_expr(sign_expr.op());
1266  out << " (_ bv0 " << op_width << "))";
1267  }
1268  else
1270  false,
1271  "sign should not be applied to unsupported type",
1272  sign_expr.type().id_string());
1273  }
1274  else if(expr.id()==ID_if)
1275  {
1276  const if_exprt &if_expr = to_if_expr(expr);
1277 
1278  out << "(ite ";
1279  convert_expr(if_expr.cond());
1280  out << " ";
1281  convert_expr(if_expr.true_case());
1282  out << " ";
1283  convert_expr(if_expr.false_case());
1284  out << ")";
1285  }
1286  else if(expr.id()==ID_and ||
1287  expr.id()==ID_or ||
1288  expr.id()==ID_xor)
1289  {
1291  expr.type().id() == ID_bool,
1292  "logical and, or, and xor expressions should have Boolean type");
1294  expr.operands().size() >= 2,
1295  "logical and, or, and xor expressions should have at least two operands");
1296 
1297  out << "(" << expr.id();
1298  forall_operands(it, expr)
1299  {
1300  out << " ";
1301  convert_expr(*it);
1302  }
1303  out << ")";
1304  }
1305  else if(expr.id()==ID_implies)
1306  {
1307  const implies_exprt &implies_expr = to_implies_expr(expr);
1308 
1310  implies_expr.type().id() == ID_bool,
1311  "implies expression should have Boolean type");
1312 
1313  out << "(=> ";
1314  convert_expr(implies_expr.op0());
1315  out << " ";
1316  convert_expr(implies_expr.op1());
1317  out << ")";
1318  }
1319  else if(expr.id()==ID_not)
1320  {
1321  const not_exprt &not_expr = to_not_expr(expr);
1322 
1324  not_expr.type().id() == ID_bool,
1325  "not expression should have Boolean type");
1326 
1327  out << "(not ";
1328  convert_expr(not_expr.op());
1329  out << ")";
1330  }
1331  else if(expr.id() == ID_equal)
1332  {
1333  const equal_exprt &equal_expr = to_equal_expr(expr);
1334 
1336  equal_expr.op0().type() == equal_expr.op1().type(),
1337  "operands of equal expression shall have same type");
1338 
1339  out << "(= ";
1340  convert_expr(equal_expr.op0());
1341  out << " ";
1342  convert_expr(equal_expr.op1());
1343  out << ")";
1344  }
1345  else if(expr.id() == ID_notequal)
1346  {
1347  const notequal_exprt &notequal_expr = to_notequal_expr(expr);
1348 
1350  notequal_expr.op0().type() == notequal_expr.op1().type(),
1351  "operands of not equal expression shall have same type");
1352 
1353  out << "(not (= ";
1354  convert_expr(notequal_expr.op0());
1355  out << " ";
1356  convert_expr(notequal_expr.op1());
1357  out << "))";
1358  }
1359  else if(expr.id()==ID_ieee_float_equal ||
1360  expr.id()==ID_ieee_float_notequal)
1361  {
1362  // These are not the same as (= A B)
1363  // because of NaN and negative zero.
1364  const auto &rel_expr = to_binary_relation_expr(expr);
1365 
1367  rel_expr.lhs().type() == rel_expr.rhs().type(),
1368  "operands of float equal and not equal expressions shall have same type");
1369 
1370  // The FPA theory properly treats NaN and negative zero.
1371  if(use_FPA_theory)
1372  {
1373  if(rel_expr.id() == ID_ieee_float_notequal)
1374  out << "(not ";
1375 
1376  out << "(fp.eq ";
1377  convert_expr(rel_expr.lhs());
1378  out << " ";
1379  convert_expr(rel_expr.rhs());
1380  out << ")";
1381 
1382  if(rel_expr.id() == ID_ieee_float_notequal)
1383  out << ")";
1384  }
1385  else
1386  convert_floatbv(expr);
1387  }
1388  else if(expr.id()==ID_le ||
1389  expr.id()==ID_lt ||
1390  expr.id()==ID_ge ||
1391  expr.id()==ID_gt)
1392  {
1394  }
1395  else if(expr.id()==ID_plus)
1396  {
1397  convert_plus(to_plus_expr(expr));
1398  }
1399  else if(expr.id()==ID_floatbv_plus)
1400  {
1402  }
1403  else if(expr.id()==ID_minus)
1404  {
1406  }
1407  else if(expr.id()==ID_floatbv_minus)
1408  {
1410  }
1411  else if(expr.id()==ID_div)
1412  {
1413  convert_div(to_div_expr(expr));
1414  }
1415  else if(expr.id()==ID_floatbv_div)
1416  {
1418  }
1419  else if(expr.id()==ID_mod)
1420  {
1421  convert_mod(to_mod_expr(expr));
1422  }
1423  else if(expr.id() == ID_euclidean_mod)
1424  {
1426  }
1427  else if(expr.id()==ID_mult)
1428  {
1429  convert_mult(to_mult_expr(expr));
1430  }
1431  else if(expr.id()==ID_floatbv_mult)
1432  {
1434  }
1435  else if(expr.id() == ID_floatbv_rem)
1436  {
1438  }
1439  else if(expr.id()==ID_address_of)
1440  {
1441  const address_of_exprt &address_of_expr = to_address_of_expr(expr);
1443  address_of_expr.object(), to_pointer_type(address_of_expr.type()));
1444  }
1445  else if(expr.id() == ID_array_of)
1446  {
1447  const auto &array_of_expr = to_array_of_expr(expr);
1448 
1450  array_of_expr.type().id() == ID_array,
1451  "array of expression shall have array type");
1452 
1453  if(use_as_const)
1454  {
1455  out << "((as const ";
1456  convert_type(array_of_expr.type());
1457  out << ") ";
1458  convert_expr(array_of_expr.what());
1459  out << ")";
1460  }
1461  else
1462  {
1463  defined_expressionst::const_iterator it =
1464  defined_expressions.find(array_of_expr);
1465  CHECK_RETURN(it != defined_expressions.end());
1466  out << it->second;
1467  }
1468  }
1469  else if(expr.id() == ID_array_comprehension)
1470  {
1471  const auto &array_comprehension = to_array_comprehension_expr(expr);
1472 
1474  array_comprehension.type().id() == ID_array,
1475  "array_comprehension expression shall have array type");
1476 
1478  {
1479  out << "(lambda ((";
1480  convert_expr(array_comprehension.arg());
1481  out << " ";
1482  convert_type(array_comprehension.type().size().type());
1483  out << ")) ";
1484  convert_expr(array_comprehension.body());
1485  out << ")";
1486  }
1487  else
1488  {
1489  const auto &it = defined_expressions.find(array_comprehension);
1490  CHECK_RETURN(it != defined_expressions.end());
1491  out << it->second;
1492  }
1493  }
1494  else if(expr.id()==ID_index)
1495  {
1497  }
1498  else if(expr.id()==ID_ashr ||
1499  expr.id()==ID_lshr ||
1500  expr.id()==ID_shl)
1501  {
1502  const shift_exprt &shift_expr = to_shift_expr(expr);
1503  const typet &type = shift_expr.type();
1504 
1505  if(type.id()==ID_unsignedbv ||
1506  type.id()==ID_signedbv ||
1507  type.id()==ID_bv)
1508  {
1509  if(shift_expr.id() == ID_ashr)
1510  out << "(bvashr ";
1511  else if(shift_expr.id() == ID_lshr)
1512  out << "(bvlshr ";
1513  else if(shift_expr.id() == ID_shl)
1514  out << "(bvshl ";
1515  else
1516  UNREACHABLE;
1517 
1518  convert_expr(shift_expr.op());
1519  out << " ";
1520 
1521  // SMT2 requires the shift distance to have the same width as
1522  // the value that is shifted -- odd!
1523 
1524  if(shift_expr.distance().type().id() == ID_integer)
1525  {
1526  const mp_integer i =
1527  numeric_cast_v<mp_integer>(to_constant_expr(shift_expr.distance()));
1528 
1529  // shift distance must be bit vector
1530  std::size_t width_op0 = boolbv_width(shift_expr.op().type());
1531  exprt tmp=from_integer(i, unsignedbv_typet(width_op0));
1532  convert_expr(tmp);
1533  }
1534  else if(
1535  shift_expr.distance().type().id() == ID_signedbv ||
1536  shift_expr.distance().type().id() == ID_unsignedbv ||
1537  shift_expr.distance().type().id() == ID_c_enum ||
1538  shift_expr.distance().type().id() == ID_c_bool)
1539  {
1540  std::size_t width_op0 = boolbv_width(shift_expr.op().type());
1541  std::size_t width_op1 = boolbv_width(shift_expr.distance().type());
1542 
1543  if(width_op0==width_op1)
1544  convert_expr(shift_expr.distance());
1545  else if(width_op0>width_op1)
1546  {
1547  out << "((_ zero_extend " << width_op0-width_op1 << ") ";
1548  convert_expr(shift_expr.distance());
1549  out << ")"; // zero_extend
1550  }
1551  else // width_op0<width_op1
1552  {
1553  out << "((_ extract " << width_op0-1 << " 0) ";
1554  convert_expr(shift_expr.distance());
1555  out << ")"; // extract
1556  }
1557  }
1558  else
1560  "unsupported distance type for " + shift_expr.id_string() + ": " +
1561  type.id_string());
1562 
1563  out << ")"; // bv*sh
1564  }
1565  else
1567  "unsupported type for " + shift_expr.id_string() + ": " +
1568  type.id_string());
1569  }
1570  else if(expr.id() == ID_rol || expr.id() == ID_ror)
1571  {
1572  const shift_exprt &shift_expr = to_shift_expr(expr);
1573  const typet &type = shift_expr.type();
1574 
1575  if(
1576  type.id() == ID_unsignedbv || type.id() == ID_signedbv ||
1577  type.id() == ID_bv)
1578  {
1579  // SMT-LIB offers rotate_left and rotate_right, but these require a
1580  // constant distance.
1581  if(shift_expr.id() == ID_rol)
1582  out << "((_ rotate_left";
1583  else if(shift_expr.id() == ID_ror)
1584  out << "((_ rotate_right";
1585  else
1586  UNREACHABLE;
1587 
1588  out << ' ';
1589 
1590  auto distance_int_op = numeric_cast<mp_integer>(shift_expr.distance());
1591 
1592  if(distance_int_op.has_value())
1593  {
1594  out << distance_int_op.value();
1595  }
1596  else
1598  "distance type for " + shift_expr.id_string() + "must be constant");
1599 
1600  out << ") ";
1601  convert_expr(shift_expr.op());
1602 
1603  out << ")"; // rotate_*
1604  }
1605  else
1607  "unsupported type for " + shift_expr.id_string() + ": " +
1608  type.id_string());
1609  }
1610  else if(expr.id() == ID_named_term)
1611  {
1612  const auto &named_term_expr = to_named_term_expr(expr);
1613  out << "(! ";
1614  convert(named_term_expr.value());
1615  out << " :named "
1616  << convert_identifier(named_term_expr.symbol().get_identifier()) << ')';
1617  }
1618  else if(expr.id()==ID_with)
1619  {
1620  convert_with(to_with_expr(expr));
1621  }
1622  else if(expr.id()==ID_update)
1623  {
1624  convert_update(expr);
1625  }
1626  else if(expr.id()==ID_member)
1627  {
1629  }
1630  else if(expr.id()==ID_pointer_offset)
1631  {
1632  const auto &op = to_pointer_offset_expr(expr).pointer();
1633 
1635  op.type().id() == ID_pointer,
1636  "operand of pointer offset expression shall be of pointer type");
1637 
1638  std::size_t offset_bits =
1640  std::size_t result_width=boolbv_width(expr.type());
1641 
1642  // max extract width
1643  if(offset_bits>result_width)
1644  offset_bits=result_width;
1645 
1646  // too few bits?
1647  if(result_width>offset_bits)
1648  out << "((_ zero_extend " << result_width-offset_bits << ") ";
1649 
1650  out << "((_ extract " << offset_bits-1 << " 0) ";
1651  convert_expr(op);
1652  out << ")";
1653 
1654  if(result_width>offset_bits)
1655  out << ")"; // zero_extend
1656  }
1657  else if(expr.id()==ID_pointer_object)
1658  {
1659  const auto &op = to_pointer_object_expr(expr).pointer();
1660 
1662  op.type().id() == ID_pointer,
1663  "pointer object expressions should be of pointer type");
1664 
1665  std::size_t ext=boolbv_width(expr.type())-config.bv_encoding.object_bits;
1666  std::size_t pointer_width = boolbv_width(op.type());
1667 
1668  if(ext>0)
1669  out << "((_ zero_extend " << ext << ") ";
1670 
1671  out << "((_ extract "
1672  << pointer_width-1 << " "
1673  << pointer_width-config.bv_encoding.object_bits << ") ";
1674  convert_expr(op);
1675  out << ")";
1676 
1677  if(ext>0)
1678  out << ")"; // zero_extend
1679  }
1680  else if(expr.id() == ID_is_dynamic_object)
1681  {
1683  }
1684  else if(expr.id() == ID_is_invalid_pointer)
1685  {
1686  const auto &op = to_unary_expr(expr).op();
1687  std::size_t pointer_width = boolbv_width(op.type());
1688  out << "(= ((_ extract "
1689  << pointer_width-1 << " "
1690  << pointer_width-config.bv_encoding.object_bits << ") ";
1691  convert_expr(op);
1692  out << ") (_ bv" << pointer_logic.get_invalid_object()
1693  << " " << config.bv_encoding.object_bits << "))";
1694  }
1695  else if(expr.id()==ID_string_constant)
1696  {
1697  defined_expressionst::const_iterator it=defined_expressions.find(expr);
1698  CHECK_RETURN(it != defined_expressions.end());
1699  out << it->second;
1700  }
1701  else if(expr.id()==ID_extractbit)
1702  {
1703  const extractbit_exprt &extractbit_expr = to_extractbit_expr(expr);
1704 
1705  if(extractbit_expr.index().is_constant())
1706  {
1707  const mp_integer i =
1708  numeric_cast_v<mp_integer>(to_constant_expr(extractbit_expr.index()));
1709 
1710  out << "(= ((_ extract " << i << " " << i << ") ";
1711  flatten2bv(extractbit_expr.src());
1712  out << ") #b1)";
1713  }
1714  else
1715  {
1716  out << "(= ((_ extract 0 0) ";
1717  // the arguments of the shift need to have the same width
1718  out << "(bvlshr ";
1719  flatten2bv(extractbit_expr.src());
1720  typecast_exprt tmp(extractbit_expr.index(), extractbit_expr.src().type());
1721  convert_expr(tmp);
1722  out << ")) bin1)"; // bvlshr, extract, =
1723  }
1724  }
1725  else if(expr.id()==ID_extractbits)
1726  {
1727  const extractbits_exprt &extractbits_expr = to_extractbits_expr(expr);
1728 
1729  if(
1730  extractbits_expr.upper().is_constant() &&
1731  extractbits_expr.lower().is_constant())
1732  {
1733  mp_integer op1_i =
1734  numeric_cast_v<mp_integer>(to_constant_expr(extractbits_expr.upper()));
1735  mp_integer op2_i =
1736  numeric_cast_v<mp_integer>(to_constant_expr(extractbits_expr.lower()));
1737 
1738  if(op2_i>op1_i)
1739  std::swap(op1_i, op2_i);
1740 
1741  // now op1_i>=op2_i
1742 
1743  out << "((_ extract " << op1_i << " " << op2_i << ") ";
1744  flatten2bv(extractbits_expr.src());
1745  out << ")";
1746  }
1747  else
1748  {
1749  #if 0
1750  out << "(= ((_ extract 0 0) ";
1751  // the arguments of the shift need to have the same width
1752  out << "(bvlshr ";
1753  convert_expr(expr.op0());
1754  typecast_exprt tmp(expr.op0().type());
1755  tmp.op0()=expr.op1();
1756  convert_expr(tmp);
1757  out << ")) bin1)"; // bvlshr, extract, =
1758  #endif
1759  SMT2_TODO("smt2: extractbits with non-constant index");
1760  }
1761  }
1762  else if(expr.id()==ID_replication)
1763  {
1764  const replication_exprt &replication_expr = to_replication_expr(expr);
1765 
1766  mp_integer times = numeric_cast_v<mp_integer>(replication_expr.times());
1767 
1768  out << "((_ repeat " << times << ") ";
1769  flatten2bv(replication_expr.op());
1770  out << ")";
1771  }
1772  else if(expr.id()==ID_byte_extract_little_endian ||
1773  expr.id()==ID_byte_extract_big_endian)
1774  {
1775  INVARIANT(
1776  false, "byte_extract ops should be lowered in prepare_for_convert_expr");
1777  }
1778  else if(expr.id()==ID_byte_update_little_endian ||
1779  expr.id()==ID_byte_update_big_endian)
1780  {
1781  INVARIANT(
1782  false, "byte_update ops should be lowered in prepare_for_convert_expr");
1783  }
1784  else if(expr.id()==ID_abs)
1785  {
1786  const abs_exprt &abs_expr = to_abs_expr(expr);
1787 
1788  const typet &type = abs_expr.type();
1789 
1790  if(type.id()==ID_signedbv)
1791  {
1792  std::size_t result_width = to_signedbv_type(type).get_width();
1793 
1794  out << "(ite (bvslt ";
1795  convert_expr(abs_expr.op());
1796  out << " (_ bv0 " << result_width << ")) ";
1797  out << "(bvneg ";
1798  convert_expr(abs_expr.op());
1799  out << ") ";
1800  convert_expr(abs_expr.op());
1801  out << ")";
1802  }
1803  else if(type.id()==ID_fixedbv)
1804  {
1805  std::size_t result_width=to_fixedbv_type(type).get_width();
1806 
1807  out << "(ite (bvslt ";
1808  convert_expr(abs_expr.op());
1809  out << " (_ bv0 " << result_width << ")) ";
1810  out << "(bvneg ";
1811  convert_expr(abs_expr.op());
1812  out << ") ";
1813  convert_expr(abs_expr.op());
1814  out << ")";
1815  }
1816  else if(type.id()==ID_floatbv)
1817  {
1818  if(use_FPA_theory)
1819  {
1820  out << "(fp.abs ";
1821  convert_expr(abs_expr.op());
1822  out << ")";
1823  }
1824  else
1825  convert_floatbv(abs_expr);
1826  }
1827  else
1828  UNREACHABLE;
1829  }
1830  else if(expr.id()==ID_isnan)
1831  {
1832  const isnan_exprt &isnan_expr = to_isnan_expr(expr);
1833 
1834  const typet &op_type = isnan_expr.op().type();
1835 
1836  if(op_type.id()==ID_fixedbv)
1837  out << "false";
1838  else if(op_type.id()==ID_floatbv)
1839  {
1840  if(use_FPA_theory)
1841  {
1842  out << "(fp.isNaN ";
1843  convert_expr(isnan_expr.op());
1844  out << ")";
1845  }
1846  else
1847  convert_floatbv(isnan_expr);
1848  }
1849  else
1850  UNREACHABLE;
1851  }
1852  else if(expr.id()==ID_isfinite)
1853  {
1854  const isfinite_exprt &isfinite_expr = to_isfinite_expr(expr);
1855 
1856  const typet &op_type = isfinite_expr.op().type();
1857 
1858  if(op_type.id()==ID_fixedbv)
1859  out << "true";
1860  else if(op_type.id()==ID_floatbv)
1861  {
1862  if(use_FPA_theory)
1863  {
1864  out << "(and ";
1865 
1866  out << "(not (fp.isNaN ";
1867  convert_expr(isfinite_expr.op());
1868  out << "))";
1869 
1870  out << "(not (fp.isInfinite ";
1871  convert_expr(isfinite_expr.op());
1872  out << "))";
1873 
1874  out << ")";
1875  }
1876  else
1877  convert_floatbv(isfinite_expr);
1878  }
1879  else
1880  UNREACHABLE;
1881  }
1882  else if(expr.id()==ID_isinf)
1883  {
1884  const isinf_exprt &isinf_expr = to_isinf_expr(expr);
1885 
1886  const typet &op_type = isinf_expr.op().type();
1887 
1888  if(op_type.id()==ID_fixedbv)
1889  out << "false";
1890  else if(op_type.id()==ID_floatbv)
1891  {
1892  if(use_FPA_theory)
1893  {
1894  out << "(fp.isInfinite ";
1895  convert_expr(isinf_expr.op());
1896  out << ")";
1897  }
1898  else
1899  convert_floatbv(isinf_expr);
1900  }
1901  else
1902  UNREACHABLE;
1903  }
1904  else if(expr.id()==ID_isnormal)
1905  {
1906  const isnormal_exprt &isnormal_expr = to_isnormal_expr(expr);
1907 
1908  const typet &op_type = isnormal_expr.op().type();
1909 
1910  if(op_type.id()==ID_fixedbv)
1911  out << "true";
1912  else if(op_type.id()==ID_floatbv)
1913  {
1914  if(use_FPA_theory)
1915  {
1916  out << "(fp.isNormal ";
1917  convert_expr(isnormal_expr.op());
1918  out << ")";
1919  }
1920  else
1921  convert_floatbv(isnormal_expr);
1922  }
1923  else
1924  UNREACHABLE;
1925  }
1926  else if(
1929  expr.id() == ID_overflow_result_plus ||
1930  expr.id() == ID_overflow_result_minus)
1931  {
1932  const bool keep_result = can_cast_expr<overflow_result_exprt>(expr);
1933 
1934  const auto &op0 = to_binary_expr(expr).op0();
1935  const auto &op1 = to_binary_expr(expr).op1();
1936 
1938  keep_result || expr.type().id() == ID_bool,
1939  "overflow plus and overflow minus expressions shall be of Boolean type");
1940 
1941  bool subtract = can_cast_expr<minus_overflow_exprt>(expr) ||
1942  expr.id() == ID_overflow_result_minus;
1943  const typet &op_type = op0.type();
1944  std::size_t width=boolbv_width(op_type);
1945 
1946  if(op_type.id()==ID_signedbv)
1947  {
1948  // an overflow occurs if the top two bits of the extended sum differ
1949  out << "(let ((?sum (";
1950  out << (subtract?"bvsub":"bvadd");
1951  out << " ((_ sign_extend 1) ";
1952  convert_expr(op0);
1953  out << ")";
1954  out << " ((_ sign_extend 1) ";
1955  convert_expr(op1);
1956  out << ")))) "; // sign_extend, bvadd/sub
1957  if(keep_result)
1958  {
1959  if(use_datatypes)
1960  {
1961  const std::string &smt_typename = datatype_map.at(expr.type());
1962 
1963  // use the constructor for the Z3 datatype
1964  out << "(mk-" << smt_typename;
1965  }
1966  else
1967  out << "(concat";
1968 
1969  out << " ((_ extract " << width - 1 << " 0) ?sum) ";
1970  if(!use_datatypes)
1971  out << "(ite ";
1972  }
1973  out << "(not (= "
1974  "((_ extract " << width << " " << width << ") ?sum) "
1975  "((_ extract " << (width-1) << " " << (width-1) << ") ?sum)";
1976  out << "))"; // =, not
1977  if(keep_result)
1978  {
1979  if(!use_datatypes)
1980  out << " #b1 #b0)";
1981  out << ")"; // concat
1982  }
1983  out << ")"; // let
1984  }
1985  else if(op_type.id()==ID_unsignedbv ||
1986  op_type.id()==ID_pointer)
1987  {
1988  // overflow is simply carry-out
1989  out << "(let ((?sum (" << (subtract ? "bvsub" : "bvadd");
1990  out << " ((_ zero_extend 1) ";
1991  convert_expr(op0);
1992  out << ")";
1993  out << " ((_ zero_extend 1) ";
1994  convert_expr(op1);
1995  out << "))))"; // zero_extend, bvsub/bvadd
1996  if(keep_result && !use_datatypes)
1997  out << " ?sum";
1998  else
1999  {
2000  if(keep_result && use_datatypes)
2001  {
2002  const std::string &smt_typename = datatype_map.at(expr.type());
2003 
2004  // use the constructor for the Z3 datatype
2005  out << "(mk-" << smt_typename;
2006  out << " ((_ extract " << width - 1 << " 0) ?sum) ";
2007  }
2008 
2009  out << "(= ";
2010  out << "((_ extract " << width << " " << width << ") ?sum)";
2011  out << "#b1)"; // =
2012 
2013  if(keep_result && use_datatypes)
2014  out << ")"; // mk
2015  }
2016  out << ")"; // let
2017  }
2018  else
2020  false,
2021  "overflow check should not be performed on unsupported type",
2022  op_type.id_string());
2023  }
2024  else if(
2026  expr.id() == ID_overflow_result_mult)
2027  {
2028  const bool keep_result = can_cast_expr<overflow_result_exprt>(expr);
2029 
2030  const auto &op0 = to_binary_expr(expr).op0();
2031  const auto &op1 = to_binary_expr(expr).op1();
2032 
2034  keep_result || expr.type().id() == ID_bool,
2035  "overflow mult expression shall be of Boolean type");
2036 
2037  // No better idea than to multiply with double the bits and then compare
2038  // with max value.
2039 
2040  const typet &op_type = op0.type();
2041  std::size_t width=boolbv_width(op_type);
2042 
2043  if(op_type.id()==ID_signedbv)
2044  {
2045  out << "(let ( (prod (bvmul ((_ sign_extend " << width << ") ";
2046  convert_expr(op0);
2047  out << ") ((_ sign_extend " << width << ") ";
2048  convert_expr(op1);
2049  out << ")) )) ";
2050  if(keep_result)
2051  {
2052  if(use_datatypes)
2053  {
2054  const std::string &smt_typename = datatype_map.at(expr.type());
2055 
2056  // use the constructor for the Z3 datatype
2057  out << "(mk-" << smt_typename;
2058  }
2059  else
2060  out << "(concat";
2061 
2062  out << " ((_ extract " << width - 1 << " 0) prod) ";
2063  if(!use_datatypes)
2064  out << "(ite ";
2065  }
2066  out << "(or (bvsge prod (_ bv" << power(2, width-1) << " "
2067  << width*2 << "))";
2068  out << " (bvslt prod (bvneg (_ bv" << power(2, width - 1) << " "
2069  << width * 2 << "))))";
2070  if(keep_result)
2071  {
2072  if(!use_datatypes)
2073  out << " #b1 #b0)";
2074  out << ")"; // concat
2075  }
2076  out << ")";
2077  }
2078  else if(op_type.id()==ID_unsignedbv)
2079  {
2080  out << "(let ((prod (bvmul ((_ zero_extend " << width << ") ";
2081  convert_expr(op0);
2082  out << ") ((_ zero_extend " << width << ") ";
2083  convert_expr(op1);
2084  out << ")))) ";
2085  if(keep_result)
2086  {
2087  if(use_datatypes)
2088  {
2089  const std::string &smt_typename = datatype_map.at(expr.type());
2090 
2091  // use the constructor for the Z3 datatype
2092  out << "(mk-" << smt_typename;
2093  }
2094  else
2095  out << "(concat";
2096 
2097  out << " ((_ extract " << width - 1 << " 0) prod) ";
2098  if(!use_datatypes)
2099  out << "(ite ";
2100  }
2101  out << "(bvuge prod (_ bv" << power(2, width) << " " << width * 2 << "))";
2102  if(keep_result)
2103  {
2104  if(!use_datatypes)
2105  out << " #b1 #b0)";
2106  out << ")"; // concat
2107  }
2108  out << ")";
2109  }
2110  else
2112  false,
2113  "overflow check should not be performed on unsupported type",
2114  op_type.id_string());
2115  }
2116  else if(expr.id()==ID_array)
2117  {
2118  defined_expressionst::const_iterator it=defined_expressions.find(expr);
2119  CHECK_RETURN(it != defined_expressions.end());
2120  out << it->second;
2121  }
2122  else if(expr.id()==ID_literal)
2123  {
2124  convert_literal(to_literal_expr(expr).get_literal());
2125  }
2126  else if(expr.id()==ID_forall ||
2127  expr.id()==ID_exists)
2128  {
2129  const quantifier_exprt &quantifier_expr = to_quantifier_expr(expr);
2130 
2132  // NOLINTNEXTLINE(readability/throw)
2133  throw "MathSAT does not support quantifiers";
2134 
2135  if(quantifier_expr.id() == ID_forall)
2136  out << "(forall ";
2137  else if(quantifier_expr.id() == ID_exists)
2138  out << "(exists ";
2139 
2140  out << '(';
2141  bool first = true;
2142  for(const auto &bound : quantifier_expr.variables())
2143  {
2144  if(first)
2145  first = false;
2146  else
2147  out << ' ';
2148  out << '(';
2149  convert_expr(bound);
2150  out << ' ';
2151  convert_type(bound.type());
2152  out << ')';
2153  }
2154  out << ") ";
2155 
2156  convert_expr(quantifier_expr.where());
2157 
2158  out << ')';
2159  }
2160  else if(expr.id()==ID_vector)
2161  {
2162  const vector_exprt &vector_expr = to_vector_expr(expr);
2163  const vector_typet &vector_type = to_vector_type(vector_expr.type());
2164 
2165  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
2166 
2168  size == vector_expr.operands().size(),
2169  "size indicated by type shall be equal to the number of operands");
2170 
2171  if(use_datatypes)
2172  {
2173  const std::string &smt_typename = datatype_map.at(vector_type);
2174 
2175  out << "(mk-" << smt_typename;
2176  }
2177  else
2178  out << "(concat";
2179 
2180  // build component-by-component
2181  forall_operands(it, vector_expr)
2182  {
2183  out << " ";
2184  convert_expr(*it);
2185  }
2186 
2187  out << ")"; // mk-... or concat
2188  }
2189  else if(
2190  const auto object_size = expr_try_dynamic_cast<object_size_exprt>(expr))
2191  {
2193  }
2194  else if(expr.id()==ID_let)
2195  {
2196  const let_exprt &let_expr=to_let_expr(expr);
2197  const auto &variables = let_expr.variables();
2198  const auto &values = let_expr.values();
2199 
2200  out << "(let (";
2201  bool first = true;
2202 
2203  for(auto &binding : make_range(variables).zip(values))
2204  {
2205  if(first)
2206  first = false;
2207  else
2208  out << ' ';
2209 
2210  out << '(';
2211  convert_expr(binding.first);
2212  out << ' ';
2213  convert_expr(binding.second);
2214  out << ')';
2215  }
2216 
2217  out << ") "; // bindings
2218 
2219  convert_expr(let_expr.where());
2220  out << ')'; // let
2221  }
2222  else if(expr.id()==ID_constraint_select_one)
2223  {
2225  "smt2_convt::convert_expr: '" + expr.id_string() +
2226  "' is not yet supported");
2227  }
2228  else if(expr.id() == ID_bswap)
2229  {
2230  const bswap_exprt &bswap_expr = to_bswap_expr(expr);
2231 
2233  bswap_expr.op().type() == bswap_expr.type(),
2234  "operand of byte swap expression shall have same type as the expression");
2235 
2236  // first 'let' the operand
2237  out << "(let ((bswap_op ";
2238  convert_expr(bswap_expr.op());
2239  out << ")) ";
2240 
2241  if(
2242  bswap_expr.type().id() == ID_signedbv ||
2243  bswap_expr.type().id() == ID_unsignedbv)
2244  {
2245  const std::size_t width =
2246  to_bitvector_type(bswap_expr.type()).get_width();
2247 
2248  const std::size_t bits_per_byte = bswap_expr.get_bits_per_byte();
2249 
2250  // width must be multiple of bytes
2252  width % bits_per_byte == 0,
2253  "bit width indicated by type of bswap expression should be a multiple "
2254  "of the number of bits per byte");
2255 
2256  const std::size_t bytes = width / bits_per_byte;
2257 
2258  if(bytes <= 1)
2259  out << "bswap_op";
2260  else
2261  {
2262  // do a parallel 'let' for each byte
2263  out << "(let (";
2264 
2265  for(std::size_t byte = 0; byte < bytes; byte++)
2266  {
2267  if(byte != 0)
2268  out << ' ';
2269  out << "(bswap_byte_" << byte << ' ';
2270  out << "((_ extract " << (byte * bits_per_byte + (bits_per_byte - 1))
2271  << " " << (byte * bits_per_byte) << ") bswap_op)";
2272  out << ')';
2273  }
2274 
2275  out << ") ";
2276 
2277  // now stitch back together with 'concat'
2278  out << "(concat";
2279 
2280  for(std::size_t byte = 0; byte < bytes; byte++)
2281  out << " bswap_byte_" << byte;
2282 
2283  out << ')'; // concat
2284  out << ')'; // let bswap_byte_*
2285  }
2286  }
2287  else
2288  UNEXPECTEDCASE("bswap must get bitvector operand");
2289 
2290  out << ')'; // let bswap_op
2291  }
2292  else if(expr.id() == ID_popcount)
2293  {
2294  convert_expr(simplify_expr(to_popcount_expr(expr).lower(), ns));
2295  }
2296  else if(expr.id() == ID_count_leading_zeros)
2297  {
2299  }
2300  else if(expr.id() == ID_count_trailing_zeros)
2301  {
2303  }
2304  else if(expr.id() == ID_find_first_set)
2305  {
2307  }
2308  else if(expr.id() == ID_empty_union)
2309  {
2310  out << "()";
2311  }
2312  else if(expr.id() == ID_bitreverse)
2313  {
2315  }
2316  else if(expr.id() == ID_function_application)
2317  {
2318  const auto &function_application_expr = to_function_application_expr(expr);
2319  // do not use parentheses if there function is a constant
2320  if(function_application_expr.arguments().empty())
2321  {
2322  convert_expr(function_application_expr.function());
2323  }
2324  else
2325  {
2326  out << '(';
2327  convert_expr(function_application_expr.function());
2328  for(auto &op : function_application_expr.arguments())
2329  {
2330  out << ' ';
2331  convert_expr(op);
2332  }
2333  out << ')';
2334  }
2335  }
2336  else
2338  false,
2339  "smt2_convt::convert_expr should not be applied to unsupported type",
2340  expr.id_string());
2341 }
2342 
2344 {
2345  const exprt &src = expr.op();
2346 
2347  typet dest_type = expr.type();
2348  if(dest_type.id()==ID_c_enum_tag)
2349  dest_type=ns.follow_tag(to_c_enum_tag_type(dest_type));
2350 
2351  typet src_type = src.type();
2352  if(src_type.id()==ID_c_enum_tag)
2353  src_type=ns.follow_tag(to_c_enum_tag_type(src_type));
2354 
2355  if(dest_type.id()==ID_bool)
2356  {
2357  // this is comparison with zero
2358  if(src_type.id()==ID_signedbv ||
2359  src_type.id()==ID_unsignedbv ||
2360  src_type.id()==ID_c_bool ||
2361  src_type.id()==ID_fixedbv ||
2362  src_type.id()==ID_pointer ||
2363  src_type.id()==ID_integer)
2364  {
2365  out << "(not (= ";
2366  convert_expr(src);
2367  out << " ";
2368  convert_expr(from_integer(0, src_type));
2369  out << "))";
2370  }
2371  else if(src_type.id()==ID_floatbv)
2372  {
2373  if(use_FPA_theory)
2374  {
2375  out << "(not (fp.isZero ";
2376  convert_expr(src);
2377  out << "))";
2378  }
2379  else
2380  convert_floatbv(expr);
2381  }
2382  else
2383  {
2384  UNEXPECTEDCASE("TODO typecast1 "+src_type.id_string()+" -> bool");
2385  }
2386  }
2387  else if(dest_type.id()==ID_c_bool)
2388  {
2389  std::size_t to_width=boolbv_width(dest_type);
2390  out << "(ite ";
2391  out << "(not (= ";
2392  convert_expr(src);
2393  out << " ";
2394  convert_expr(from_integer(0, src_type));
2395  out << "))"; // not, =
2396  out << " (_ bv1 " << to_width << ")";
2397  out << " (_ bv0 " << to_width << ")";
2398  out << ")"; // ite
2399  }
2400  else if(dest_type.id()==ID_signedbv ||
2401  dest_type.id()==ID_unsignedbv ||
2402  dest_type.id()==ID_c_enum ||
2403  dest_type.id()==ID_bv)
2404  {
2405  std::size_t to_width=boolbv_width(dest_type);
2406 
2407  if(src_type.id()==ID_signedbv || // from signedbv
2408  src_type.id()==ID_unsignedbv || // from unsigedbv
2409  src_type.id()==ID_c_bool ||
2410  src_type.id()==ID_c_enum ||
2411  src_type.id()==ID_bv)
2412  {
2413  std::size_t from_width=boolbv_width(src_type);
2414 
2415  if(from_width==to_width)
2416  convert_expr(src); // ignore
2417  else if(from_width<to_width) // extend
2418  {
2419  if(src_type.id()==ID_signedbv)
2420  out << "((_ sign_extend ";
2421  else
2422  out << "((_ zero_extend ";
2423 
2424  out << (to_width-from_width)
2425  << ") "; // ind
2426  convert_expr(src);
2427  out << ")";
2428  }
2429  else // chop off extra bits
2430  {
2431  out << "((_ extract " << (to_width-1) << " 0) ";
2432  convert_expr(src);
2433  out << ")";
2434  }
2435  }
2436  else if(src_type.id()==ID_fixedbv) // from fixedbv to int
2437  {
2438  const fixedbv_typet &fixedbv_type=to_fixedbv_type(src_type);
2439 
2440  std::size_t from_width=fixedbv_type.get_width();
2441  std::size_t from_integer_bits=fixedbv_type.get_integer_bits();
2442  std::size_t from_fraction_bits=fixedbv_type.get_fraction_bits();
2443 
2444  // we might need to round up in case of negative numbers
2445  // e.g., (int)(-1.00001)==1
2446 
2447  out << "(let ((?tcop ";
2448  convert_expr(src);
2449  out << ")) ";
2450 
2451  out << "(bvadd ";
2452 
2453  if(to_width>from_integer_bits)
2454  {
2455  out << "((_ sign_extend "
2456  << (to_width-from_integer_bits) << ") ";
2457  out << "((_ extract " << (from_width-1) << " "
2458  << from_fraction_bits << ") ";
2459  convert_expr(src);
2460  out << "))";
2461  }
2462  else
2463  {
2464  out << "((_ extract " << (from_fraction_bits+to_width-1)
2465  << " " << from_fraction_bits << ") ";
2466  convert_expr(src);
2467  out << ")";
2468  }
2469 
2470  out << " (ite (and ";
2471 
2472  // some fraction bit is not zero
2473  out << "(not (= ((_ extract " << (from_fraction_bits-1) << " 0) ?tcop) "
2474  "(_ bv0 " << from_fraction_bits << ")))";
2475 
2476  // number negative
2477  out << " (= ((_ extract " << (from_width-1) << " " << (from_width-1)
2478  << ") ?tcop) #b1)";
2479 
2480  out << ")"; // and
2481 
2482  out << " (_ bv1 " << to_width << ") (_ bv0 " << to_width << "))"; // ite
2483  out << ")"; // bvadd
2484  out << ")"; // let
2485  }
2486  else if(src_type.id()==ID_floatbv) // from floatbv to int
2487  {
2488  if(dest_type.id()==ID_bv)
2489  {
2490  // this is _NOT_ a semantic conversion, but bit-wise
2491 
2492  if(use_FPA_theory)
2493  {
2494  // This conversion is non-trivial as it requires creating a
2495  // new bit-vector variable and then asserting that it converts
2496  // to the required floating-point number.
2497  SMT2_TODO("bit-wise floatbv to bv");
2498  }
2499  else
2500  {
2501  // straight-forward if width matches
2502  convert_expr(src);
2503  }
2504  }
2505  else if(dest_type.id()==ID_signedbv)
2506  {
2507  // this should be floatbv_typecast, not typecast
2509  "typecast unexpected "+src_type.id_string()+" -> "+
2510  dest_type.id_string());
2511  }
2512  else if(dest_type.id()==ID_unsignedbv)
2513  {
2514  // this should be floatbv_typecast, not typecast
2516  "typecast unexpected "+src_type.id_string()+" -> "+
2517  dest_type.id_string());
2518  }
2519  }
2520  else if(src_type.id()==ID_bool) // from boolean to int
2521  {
2522  out << "(ite ";
2523  convert_expr(src);
2524 
2525  if(dest_type.id()==ID_fixedbv)
2526  {
2527  fixedbv_spect spec(to_fixedbv_type(dest_type));
2528  out << " (concat (_ bv1 "
2529  << spec.integer_bits << ") " <<
2530  "(_ bv0 " << spec.get_fraction_bits() << ")) " <<
2531  "(_ bv0 " << spec.width << ")";
2532  }
2533  else
2534  {
2535  out << " (_ bv1 " << to_width << ")";
2536  out << " (_ bv0 " << to_width << ")";
2537  }
2538 
2539  out << ")";
2540  }
2541  else if(src_type.id()==ID_pointer) // from pointer to int
2542  {
2543  std::size_t from_width=boolbv_width(src_type);
2544 
2545  if(from_width<to_width) // extend
2546  {
2547  out << "((_ sign_extend ";
2548  out << (to_width-from_width)
2549  << ") ";
2550  convert_expr(src);
2551  out << ")";
2552  }
2553  else // chop off extra bits
2554  {
2555  out << "((_ extract " << (to_width-1) << " 0) ";
2556  convert_expr(src);
2557  out << ")";
2558  }
2559  }
2560  else if(src_type.id()==ID_integer) // from integer to bit-vector
2561  {
2562  // must be constant
2563  if(src.is_constant())
2564  {
2565  mp_integer i = numeric_cast_v<mp_integer>(to_constant_expr(src));
2566  out << "(_ bv" << i << " " << to_width << ")";
2567  }
2568  else
2569  SMT2_TODO("can't convert non-constant integer to bitvector");
2570  }
2571  else if(
2572  src_type.id() == ID_struct ||
2573  src_type.id() == ID_struct_tag) // flatten a struct to a bit-vector
2574  {
2575  if(use_datatypes)
2576  {
2577  INVARIANT(
2578  boolbv_width(src_type) == boolbv_width(dest_type),
2579  "bit vector with of source and destination type shall be equal");
2580  flatten2bv(src);
2581  }
2582  else
2583  {
2584  INVARIANT(
2585  boolbv_width(src_type) == boolbv_width(dest_type),
2586  "bit vector with of source and destination type shall be equal");
2587  convert_expr(src); // nothing else to do!
2588  }
2589  }
2590  else if(
2591  src_type.id() == ID_union ||
2592  src_type.id() == ID_union_tag) // flatten a union
2593  {
2594  INVARIANT(
2595  boolbv_width(src_type) == boolbv_width(dest_type),
2596  "bit vector with of source and destination type shall be equal");
2597  convert_expr(src); // nothing else to do!
2598  }
2599  else if(src_type.id()==ID_c_bit_field)
2600  {
2601  std::size_t from_width=boolbv_width(src_type);
2602 
2603  if(from_width==to_width)
2604  convert_expr(src); // ignore
2605  else
2606  {
2608  typecast_exprt tmp(typecast_exprt(src, t), dest_type);
2609  convert_typecast(tmp);
2610  }
2611  }
2612  else
2613  {
2614  std::ostringstream e_str;
2615  e_str << src_type.id() << " -> " << dest_type.id()
2616  << " src == " << format(src);
2617  UNEXPECTEDCASE("TODO typecast2 " + e_str.str());
2618  }
2619  }
2620  else if(dest_type.id()==ID_fixedbv) // to fixedbv
2621  {
2622  const fixedbv_typet &fixedbv_type=to_fixedbv_type(dest_type);
2623  std::size_t to_fraction_bits=fixedbv_type.get_fraction_bits();
2624  std::size_t to_integer_bits=fixedbv_type.get_integer_bits();
2625 
2626  if(src_type.id()==ID_unsignedbv ||
2627  src_type.id()==ID_signedbv ||
2628  src_type.id()==ID_c_enum)
2629  {
2630  // integer to fixedbv
2631 
2632  std::size_t from_width=to_bitvector_type(src_type).get_width();
2633  out << "(concat ";
2634 
2635  if(from_width==to_integer_bits)
2636  convert_expr(src);
2637  else if(from_width>to_integer_bits)
2638  {
2639  // too many integer bits
2640  out << "((_ extract " << (to_integer_bits-1) << " 0) ";
2641  convert_expr(src);
2642  out << ")";
2643  }
2644  else
2645  {
2646  // too few integer bits
2647  INVARIANT(
2648  from_width < to_integer_bits,
2649  "from_width should be smaller than to_integer_bits as other case "
2650  "have been handled above");
2651  if(dest_type.id()==ID_unsignedbv)
2652  {
2653  out << "(_ zero_extend "
2654  << (to_integer_bits-from_width) << ") ";
2655  convert_expr(src);
2656  out << ")";
2657  }
2658  else
2659  {
2660  out << "((_ sign_extend "
2661  << (to_integer_bits-from_width) << ") ";
2662  convert_expr(src);
2663  out << ")";
2664  }
2665  }
2666 
2667  out << "(_ bv0 " << to_fraction_bits << ")";
2668  out << ")"; // concat
2669  }
2670  else if(src_type.id()==ID_bool) // bool to fixedbv
2671  {
2672  out << "(concat (concat"
2673  << " (_ bv0 " << (to_integer_bits-1) << ") ";
2674  flatten2bv(src); // produces #b0 or #b1
2675  out << ") (_ bv0 "
2676  << to_fraction_bits
2677  << "))";
2678  }
2679  else if(src_type.id()==ID_fixedbv) // fixedbv to fixedbv
2680  {
2681  const fixedbv_typet &from_fixedbv_type=to_fixedbv_type(src_type);
2682  std::size_t from_fraction_bits=from_fixedbv_type.get_fraction_bits();
2683  std::size_t from_integer_bits=from_fixedbv_type.get_integer_bits();
2684  std::size_t from_width=from_fixedbv_type.get_width();
2685 
2686  out << "(let ((?tcop ";
2687  convert_expr(src);
2688  out << ")) ";
2689 
2690  out << "(concat ";
2691 
2692  if(to_integer_bits<=from_integer_bits)
2693  {
2694  out << "((_ extract "
2695  << (from_fraction_bits+to_integer_bits-1) << " "
2696  << from_fraction_bits
2697  << ") ?tcop)";
2698  }
2699  else
2700  {
2701  INVARIANT(
2702  to_integer_bits > from_integer_bits,
2703  "to_integer_bits should be greater than from_integer_bits as the"
2704  "other case has been handled above");
2705  out << "((_ sign_extend "
2706  << (to_integer_bits-from_integer_bits)
2707  << ") ((_ extract "
2708  << (from_width-1) << " "
2709  << from_fraction_bits
2710  << ") ?tcop))";
2711  }
2712 
2713  out << " ";
2714 
2715  if(to_fraction_bits<=from_fraction_bits)
2716  {
2717  out << "((_ extract "
2718  << (from_fraction_bits-1) << " "
2719  << (from_fraction_bits-to_fraction_bits)
2720  << ") ?tcop)";
2721  }
2722  else
2723  {
2724  INVARIANT(
2725  to_fraction_bits > from_fraction_bits,
2726  "to_fraction_bits should be greater than from_fraction_bits as the"
2727  "other case has been handled above");
2728  out << "(concat ((_ extract "
2729  << (from_fraction_bits-1) << " 0) ";
2730  convert_expr(src);
2731  out << ")"
2732  << " (_ bv0 " << to_fraction_bits-from_fraction_bits
2733  << "))";
2734  }
2735 
2736  out << "))"; // concat, let
2737  }
2738  else
2739  UNEXPECTEDCASE("unexpected typecast to fixedbv");
2740  }
2741  else if(dest_type.id()==ID_pointer)
2742  {
2743  std::size_t to_width=boolbv_width(dest_type);
2744 
2745  if(src_type.id()==ID_pointer) // pointer to pointer
2746  {
2747  // this just passes through
2748  convert_expr(src);
2749  }
2750  else if(
2751  src_type.id() == ID_unsignedbv || src_type.id() == ID_signedbv ||
2752  src_type.id() == ID_bv)
2753  {
2754  // integer to pointer
2755 
2756  std::size_t from_width=boolbv_width(src_type);
2757 
2758  if(from_width==to_width)
2759  convert_expr(src);
2760  else if(from_width<to_width)
2761  {
2762  out << "((_ sign_extend "
2763  << (to_width-from_width)
2764  << ") ";
2765  convert_expr(src);
2766  out << ")"; // sign_extend
2767  }
2768  else // from_width>to_width
2769  {
2770  out << "((_ extract " << to_width << " 0) ";
2771  convert_expr(src);
2772  out << ")"; // extract
2773  }
2774  }
2775  else
2776  UNEXPECTEDCASE("TODO typecast3 "+src_type.id_string()+" -> pointer");
2777  }
2778  else if(dest_type.id()==ID_range)
2779  {
2780  SMT2_TODO("range typecast");
2781  }
2782  else if(dest_type.id()==ID_floatbv)
2783  {
2784  // Typecast from integer to floating-point should have be been
2785  // converted to ID_floatbv_typecast during symbolic execution,
2786  // adding the rounding mode. See
2787  // smt2_convt::convert_floatbv_typecast.
2788  // The exception is bool and c_bool to float.
2789  const auto &dest_floatbv_type = to_floatbv_type(dest_type);
2790 
2791  if(src_type.id()==ID_bool)
2792  {
2793  constant_exprt val(irep_idt(), dest_type);
2794 
2795  ieee_floatt a(dest_floatbv_type);
2796 
2797  mp_integer significand;
2798  mp_integer exponent;
2799 
2800  out << "(ite ";
2801  convert_expr(src);
2802  out << " ";
2803 
2804  significand = 1;
2805  exponent = 0;
2806  a.build(significand, exponent);
2807  val.set_value(integer2bvrep(a.pack(), a.spec.width()));
2808 
2809  convert_constant(val);
2810  out << " ";
2811 
2812  significand = 0;
2813  exponent = 0;
2814  a.build(significand, exponent);
2815  val.set_value(integer2bvrep(a.pack(), a.spec.width()));
2816 
2817  convert_constant(val);
2818  out << ")";
2819  }
2820  else if(src_type.id()==ID_c_bool)
2821  {
2822  // turn into proper bool
2823  const typecast_exprt tmp(src, bool_typet());
2824  convert_typecast(typecast_exprt(tmp, dest_type));
2825  }
2826  else if(src_type.id() == ID_bv)
2827  {
2828  if(to_bv_type(src_type).get_width() != dest_floatbv_type.get_width())
2829  {
2830  UNEXPECTEDCASE("Typecast bv -> float with wrong width");
2831  }
2832 
2833  if(use_FPA_theory)
2834  {
2835  out << "((_ to_fp " << dest_floatbv_type.get_e() << " "
2836  << dest_floatbv_type.get_f() + 1 << ") ";
2837  convert_expr(src);
2838  out << ')';
2839  }
2840  else
2841  convert_expr(src);
2842  }
2843  else
2844  UNEXPECTEDCASE("Unknown typecast "+src_type.id_string()+" -> float");
2845  }
2846  else if(dest_type.id()==ID_integer)
2847  {
2848  if(src_type.id()==ID_bool)
2849  {
2850  out << "(ite ";
2851  convert_expr(src);
2852  out <<" 1 0)";
2853  }
2854  else
2855  UNEXPECTEDCASE("Unknown typecast "+src_type.id_string()+" -> integer");
2856  }
2857  else if(dest_type.id()==ID_c_bit_field)
2858  {
2859  std::size_t from_width=boolbv_width(src_type);
2860  std::size_t to_width=boolbv_width(dest_type);
2861 
2862  if(from_width==to_width)
2863  convert_expr(src); // ignore
2864  else
2865  {
2867  typecast_exprt tmp(typecast_exprt(src, t), dest_type);
2868  convert_typecast(tmp);
2869  }
2870  }
2871  else
2873  "TODO typecast8 "+src_type.id_string()+" -> "+dest_type.id_string());
2874 }
2875 
2877 {
2878  const exprt &src=expr.op();
2879  // const exprt &rounding_mode=expr.rounding_mode();
2880  const typet &src_type=src.type();
2881  const typet &dest_type=expr.type();
2882 
2883  if(dest_type.id()==ID_floatbv)
2884  {
2885  if(src_type.id()==ID_floatbv)
2886  {
2887  // float to float
2888 
2889  /* ISO 9899:1999
2890  * 6.3.1.5 Real floating types
2891  * 1 When a float is promoted to double or long double, or a
2892  * double is promoted to long double, its value is unchanged.
2893  *
2894  * 2 When a double is demoted to float, a long double is
2895  * demoted to double or float, or a value being represented in
2896  * greater precision and range than required by its semantic
2897  * type (see 6.3.1.8) is explicitly converted to its semantic
2898  * type, if the value being converted can be represented
2899  * exactly in the new type, it is unchanged. If the value
2900  * being converted is in the range of values that can be
2901  * represented but cannot be represented exactly, the result
2902  * is either the nearest higher or nearest lower representable
2903  * value, chosen in an implementation-defined manner. If the
2904  * value being converted is outside the range of values that
2905  * can be represented, the behavior is undefined.
2906  */
2907 
2908  const floatbv_typet &dst=to_floatbv_type(dest_type);
2909 
2910  if(use_FPA_theory)
2911  {
2912  out << "((_ to_fp " << dst.get_e() << " "
2913  << dst.get_f() + 1 << ") ";
2915  out << " ";
2916  convert_expr(src);
2917  out << ")";
2918  }
2919  else
2920  convert_floatbv(expr);
2921  }
2922  else if(src_type.id()==ID_unsignedbv)
2923  {
2924  // unsigned to float
2925 
2926  /* ISO 9899:1999
2927  * 6.3.1.4 Real floating and integer
2928  * 2 When a value of integer type is converted to a real
2929  * floating type, if the value being converted can be
2930  * represented exactly in the new type, it is unchanged. If the
2931  * value being converted is in the range of values that can be
2932  * represented but cannot be represented exactly, the result is
2933  * either the nearest higher or nearest lower representable
2934  * value, chosen in an implementation-defined manner. If the
2935  * value being converted is outside the range of values that can
2936  * be represented, the behavior is undefined.
2937  */
2938 
2939  const floatbv_typet &dst=to_floatbv_type(dest_type);
2940 
2941  if(use_FPA_theory)
2942  {
2943  out << "((_ to_fp_unsigned " << dst.get_e() << " "
2944  << dst.get_f() + 1 << ") ";
2946  out << " ";
2947  convert_expr(src);
2948  out << ")";
2949  }
2950  else
2951  convert_floatbv(expr);
2952  }
2953  else if(src_type.id()==ID_signedbv)
2954  {
2955  // signed to float
2956 
2957  const floatbv_typet &dst=to_floatbv_type(dest_type);
2958 
2959  if(use_FPA_theory)
2960  {
2961  out << "((_ to_fp " << dst.get_e() << " "
2962  << dst.get_f() + 1 << ") ";
2964  out << " ";
2965  convert_expr(src);
2966  out << ")";
2967  }
2968  else
2969  convert_floatbv(expr);
2970  }
2971  else if(src_type.id()==ID_c_enum_tag)
2972  {
2973  // enum to float
2974 
2975  // We first convert to 'underlying type'
2976  floatbv_typecast_exprt tmp=expr;
2977  tmp.op() = typecast_exprt(
2978  src, ns.follow_tag(to_c_enum_tag_type(src_type)).underlying_type());
2980  }
2981  else
2983  "TODO typecast11 "+src_type.id_string()+" -> "+dest_type.id_string());
2984  }
2985  else if(dest_type.id()==ID_signedbv)
2986  {
2987  if(use_FPA_theory)
2988  {
2989  std::size_t dest_width=to_signedbv_type(dest_type).get_width();
2990  out << "((_ fp.to_sbv " << dest_width << ") ";
2992  out << " ";
2993  convert_expr(src);
2994  out << ")";
2995  }
2996  else
2997  convert_floatbv(expr);
2998  }
2999  else if(dest_type.id()==ID_unsignedbv)
3000  {
3001  if(use_FPA_theory)
3002  {
3003  std::size_t dest_width=to_unsignedbv_type(dest_type).get_width();
3004  out << "((_ fp.to_ubv " << dest_width << ") ";
3006  out << " ";
3007  convert_expr(src);
3008  out << ")";
3009  }
3010  else
3011  convert_floatbv(expr);
3012  }
3013  else
3014  {
3016  "TODO typecast12 "+src_type.id_string()+" -> "+dest_type.id_string());
3017  }
3018 }
3019 
3021 {
3022  const struct_typet &struct_type = to_struct_type(ns.follow(expr.type()));
3023 
3024  const struct_typet::componentst &components=
3025  struct_type.components();
3026 
3028  components.size() == expr.operands().size(),
3029  "number of struct components as indicated by the struct type shall be equal"
3030  "to the number of operands of the struct expression");
3031 
3032  DATA_INVARIANT(!components.empty(), "struct shall have struct components");
3033 
3034  if(use_datatypes)
3035  {
3036  const std::string &smt_typename = datatype_map.at(struct_type);
3037 
3038  // use the constructor for the Z3 datatype
3039  out << "(mk-" << smt_typename;
3040 
3041  std::size_t i=0;
3042  for(struct_typet::componentst::const_iterator
3043  it=components.begin();
3044  it!=components.end();
3045  it++, i++)
3046  {
3047  out << " ";
3048  convert_expr(expr.operands()[i]);
3049  }
3050 
3051  out << ")";
3052  }
3053  else
3054  {
3055  if(components.size()==1)
3056  convert_expr(expr.op0());
3057  else
3058  {
3059  // SMT-LIB 2 concat is binary only
3060  for(std::size_t i=components.size(); i>1; i--)
3061  {
3062  out << "(concat ";
3063 
3064  exprt op=expr.operands()[i-1];
3065 
3066  // may need to flatten array-theory arrays in there
3067  if(op.type().id() == ID_array)
3068  flatten_array(op);
3069  else if(op.type().id() == ID_bool)
3070  flatten2bv(op);
3071  else
3072  convert_expr(op);
3073 
3074  out << " ";
3075  }
3076 
3077  convert_expr(expr.op0());
3078 
3079  for(std::size_t i=1; i<components.size(); i++)
3080  out << ")";
3081  }
3082  }
3083 }
3084 
3087 {
3088  const array_typet &array_type = to_array_type(expr.type());
3089  const auto &size_expr = array_type.size();
3090  PRECONDITION(size_expr.id() == ID_constant);
3091 
3092  mp_integer size = numeric_cast_v<mp_integer>(to_constant_expr(size_expr));
3093  CHECK_RETURN_WITH_DIAGNOSTICS(size != 0, "can't convert zero-sized array");
3094 
3095  out << "(let ((?far ";
3096  convert_expr(expr);
3097  out << ")) ";
3098 
3099  for(mp_integer i=size; i!=0; --i)
3100  {
3101  if(i!=1)
3102  out << "(concat ";
3103  out << "(select ?far ";
3104  convert_expr(from_integer(i - 1, array_type.index_type()));
3105  out << ")";
3106  if(i!=1)
3107  out << " ";
3108  }
3109 
3110  // close the many parentheses
3111  for(mp_integer i=size; i>1; --i)
3112  out << ")";
3113 
3114  out << ")"; // let
3115 }
3116 
3118 {
3119  const union_typet &union_type = to_union_type(ns.follow(expr.type()));
3120  const exprt &op=expr.op();
3121 
3122  std::size_t total_width=boolbv_width(union_type);
3123 
3124  std::size_t member_width=boolbv_width(op.type());
3125 
3126  if(total_width==member_width)
3127  {
3128  flatten2bv(op);
3129  }
3130  else
3131  {
3132  // we will pad with zeros, but non-det would be better
3133  INVARIANT(
3134  total_width > member_width,
3135  "total_width should be greater than member_width as member_width can be"
3136  "at most as large as total_width and the other case has been handled "
3137  "above");
3138  out << "(concat ";
3139  out << "(_ bv0 "
3140  << (total_width-member_width) << ") ";
3141  flatten2bv(op);
3142  out << ")";
3143  }
3144 }
3145 
3147 {
3148  const typet &expr_type=expr.type();
3149 
3150  if(expr_type.id()==ID_unsignedbv ||
3151  expr_type.id()==ID_signedbv ||
3152  expr_type.id()==ID_bv ||
3153  expr_type.id()==ID_c_enum ||
3154  expr_type.id()==ID_c_enum_tag ||
3155  expr_type.id()==ID_c_bool ||
3156  expr_type.id()==ID_c_bit_field)
3157  {
3158  const std::size_t width = boolbv_width(expr_type);
3159 
3160  const mp_integer value = bvrep2integer(expr.get_value(), width, false);
3161 
3162  out << "(_ bv" << value
3163  << " " << width << ")";
3164  }
3165  else if(expr_type.id()==ID_fixedbv)
3166  {
3167  const fixedbv_spect spec(to_fixedbv_type(expr_type));
3168 
3169  const mp_integer v = bvrep2integer(expr.get_value(), spec.width, false);
3170 
3171  out << "(_ bv" << v << " " << spec.width << ")";
3172  }
3173  else if(expr_type.id()==ID_floatbv)
3174  {
3175  const floatbv_typet &floatbv_type=
3176  to_floatbv_type(expr_type);
3177 
3178  if(use_FPA_theory)
3179  {
3180  /* CBMC stores floating point literals in the most
3181  computationally useful form; biased exponents and
3182  significands including the hidden bit. Thus some encoding
3183  is needed to get to IEEE-754 style representations. */
3184 
3185  ieee_floatt v=ieee_floatt(expr);
3186  size_t e=floatbv_type.get_e();
3187  size_t f=floatbv_type.get_f()+1;
3188 
3189  /* Should be sufficient, but not currently supported by mathsat */
3190  #if 0
3191  mp_integer binary = v.pack();
3192 
3193  out << "((_ to_fp " << e << " " << f << ")"
3194  << " #b" << integer2binary(v.pack(), v.spec.width()) << ")";
3195  #endif
3196 
3197  if(v.is_NaN())
3198  {
3199  out << "(_ NaN " << e << " " << f << ")";
3200  }
3201  else if(v.is_infinity())
3202  {
3203  if(v.get_sign())
3204  out << "(_ -oo " << e << " " << f << ")";
3205  else
3206  out << "(_ +oo " << e << " " << f << ")";
3207  }
3208  else
3209  {
3210  // Zero, normal or subnormal
3211 
3212  mp_integer binary = v.pack();
3213  std::string binaryString(integer2binary(v.pack(), v.spec.width()));
3214 
3215  out << "(fp "
3216  << "#b" << binaryString.substr(0, 1) << " "
3217  << "#b" << binaryString.substr(1, e) << " "
3218  << "#b" << binaryString.substr(1+e, f-1) << ")";
3219  }
3220  }
3221  else
3222  {
3223  // produce corresponding bit-vector
3224  const ieee_float_spect spec(floatbv_type);
3225  const mp_integer v = bvrep2integer(expr.get_value(), spec.width(), false);
3226  out << "(_ bv" << v << " " << spec.width() << ")";
3227  }
3228  }
3229  else if(expr_type.id()==ID_pointer)
3230  {
3231  if(is_null_pointer(expr))
3232  {
3233  out << "(_ bv0 " << boolbv_width(expr_type)
3234  << ")";
3235  }
3236  else
3238  "unknown pointer constant: " + id2string(expr.get_value()));
3239  }
3240  else if(expr_type.id()==ID_bool)
3241  {
3242  if(expr.is_true())
3243  out << "true";
3244  else if(expr.is_false())
3245  out << "false";
3246  else
3247  UNEXPECTEDCASE("unknown Boolean constant");
3248  }
3249  else if(expr_type.id()==ID_array)
3250  {
3251  defined_expressionst::const_iterator it=defined_expressions.find(expr);
3252  CHECK_RETURN(it != defined_expressions.end());
3253  out << it->second;
3254  }
3255  else if(expr_type.id()==ID_rational)
3256  {
3257  std::string value=id2string(expr.get_value());
3258  const bool negative = has_prefix(value, "-");
3259 
3260  if(negative)
3261  out << "(- ";
3262 
3263  size_t pos=value.find("/");
3264 
3265  if(pos==std::string::npos)
3266  out << value << ".0";
3267  else
3268  {
3269  out << "(/ " << value.substr(0, pos) << ".0 "
3270  << value.substr(pos+1) << ".0)";
3271  }
3272 
3273  if(negative)
3274  out << ')';
3275  }
3276  else if(expr_type.id()==ID_integer)
3277  {
3278  const auto value = id2string(expr.get_value());
3279 
3280  // SMT2 has no negative integer literals
3281  if(has_prefix(value, "-"))
3282  out << "(- " << value.substr(1, std::string::npos) << ')';
3283  else
3284  out << value;
3285  }
3286  else
3287  UNEXPECTEDCASE("unknown constant: "+expr_type.id_string());
3288 }
3289 
3291 {
3292  if(expr.type().id() == ID_integer)
3293  {
3294  out << "(mod ";
3295  convert_expr(expr.op0());
3296  out << ' ';
3297  convert_expr(expr.op1());
3298  out << ')';
3299  }
3300  else
3302  "unsupported type for euclidean_mod: " + expr.type().id_string());
3303 }
3304 
3306 {
3307  if(expr.type().id()==ID_unsignedbv ||
3308  expr.type().id()==ID_signedbv)
3309  {
3310  if(expr.type().id()==ID_unsignedbv)
3311  out << "(bvurem ";
3312  else
3313  out << "(bvsrem ";
3314 
3315  convert_expr(expr.op0());
3316  out << " ";
3317  convert_expr(expr.op1());
3318  out << ")";
3319  }
3320  else
3321  UNEXPECTEDCASE("unsupported type for mod: "+expr.type().id_string());
3322 }
3323 
3325 {
3326  std::vector<mp_integer> dynamic_objects;
3327  pointer_logic.get_dynamic_objects(dynamic_objects);
3328 
3329  if(dynamic_objects.empty())
3330  out << "false";
3331  else
3332  {
3333  std::size_t pointer_width = boolbv_width(expr.op().type());
3334 
3335  out << "(let ((?obj ((_ extract "
3336  << pointer_width-1 << " "
3337  << pointer_width-config.bv_encoding.object_bits << ") ";
3338  convert_expr(expr.op());
3339  out << "))) ";
3340 
3341  if(dynamic_objects.size()==1)
3342  {
3343  out << "(= (_ bv" << dynamic_objects.front()
3344  << " " << config.bv_encoding.object_bits << ") ?obj)";
3345  }
3346  else
3347  {
3348  out << "(or";
3349 
3350  for(const auto &object : dynamic_objects)
3351  out << " (= (_ bv" << object
3352  << " " << config.bv_encoding.object_bits << ") ?obj)";
3353 
3354  out << ")"; // or
3355  }
3356 
3357  out << ")"; // let
3358  }
3359 }
3360 
3362 {
3363  const typet &op_type=expr.op0().type();
3364 
3365  if(op_type.id()==ID_unsignedbv ||
3366  op_type.id()==ID_bv)
3367  {
3368  out << "(";
3369  if(expr.id()==ID_le)
3370  out << "bvule";
3371  else if(expr.id()==ID_lt)
3372  out << "bvult";
3373  else if(expr.id()==ID_ge)
3374  out << "bvuge";
3375  else if(expr.id()==ID_gt)
3376  out << "bvugt";
3377 
3378  out << " ";
3379  convert_expr(expr.op0());
3380  out << " ";
3381  convert_expr(expr.op1());
3382  out << ")";
3383  }
3384  else if(op_type.id()==ID_signedbv ||
3385  op_type.id()==ID_fixedbv)
3386  {
3387  out << "(";
3388  if(expr.id()==ID_le)
3389  out << "bvsle";
3390  else if(expr.id()==ID_lt)
3391  out << "bvslt";
3392  else if(expr.id()==ID_ge)
3393  out << "bvsge";
3394  else if(expr.id()==ID_gt)
3395  out << "bvsgt";
3396 
3397  out << " ";
3398  convert_expr(expr.op0());
3399  out << " ";
3400  convert_expr(expr.op1());
3401  out << ")";
3402  }
3403  else if(op_type.id()==ID_floatbv)
3404  {
3405  if(use_FPA_theory)
3406  {
3407  out << "(";
3408  if(expr.id()==ID_le)
3409  out << "fp.leq";
3410  else if(expr.id()==ID_lt)
3411  out << "fp.lt";
3412  else if(expr.id()==ID_ge)
3413  out << "fp.geq";
3414  else if(expr.id()==ID_gt)
3415  out << "fp.gt";
3416 
3417  out << " ";
3418  convert_expr(expr.op0());
3419  out << " ";
3420  convert_expr(expr.op1());
3421  out << ")";
3422  }
3423  else
3424  convert_floatbv(expr);
3425  }
3426  else if(op_type.id()==ID_rational ||
3427  op_type.id()==ID_integer)
3428  {
3429  out << "(";
3430  out << expr.id();
3431 
3432  out << " ";
3433  convert_expr(expr.op0());
3434  out << " ";
3435  convert_expr(expr.op1());
3436  out << ")";
3437  }
3438  else if(op_type.id() == ID_pointer)
3439  {
3440  const exprt same_object = ::same_object(expr.op0(), expr.op1());
3441 
3442  out << "(and ";
3444 
3445  out << " (";
3446  if(expr.id() == ID_le)
3447  out << "bvsle";
3448  else if(expr.id() == ID_lt)
3449  out << "bvslt";
3450  else if(expr.id() == ID_ge)
3451  out << "bvsge";
3452  else if(expr.id() == ID_gt)
3453  out << "bvsgt";
3454 
3455  out << ' ';
3456  convert_expr(pointer_offset(expr.op0()));
3457  out << ' ';
3458  convert_expr(pointer_offset(expr.op1()));
3459  out << ')';
3460 
3461  out << ')';
3462  }
3463  else
3465  "unsupported type for "+expr.id_string()+": "+op_type.id_string());
3466 }
3467 
3469 {
3470  if(
3471  expr.type().id() == ID_rational || expr.type().id() == ID_integer ||
3472  expr.type().id() == ID_real)
3473  {
3474  // these are multi-ary in SMT-LIB2
3475  out << "(+";
3476 
3477  for(const auto &op : expr.operands())
3478  {
3479  out << ' ';
3480  convert_expr(op);
3481  }
3482 
3483  out << ')';
3484  }
3485  else if(
3486  expr.type().id() == ID_unsignedbv || expr.type().id() == ID_signedbv ||
3487  expr.type().id() == ID_fixedbv)
3488  {
3489  // These could be chained, i.e., need not be binary,
3490  // but at least MathSat doesn't like that.
3491  if(expr.operands().size() == 2)
3492  {
3493  out << "(bvadd ";
3494  convert_expr(expr.op0());
3495  out << " ";
3496  convert_expr(expr.op1());
3497  out << ")";
3498  }
3499  else
3500  {
3502  }
3503  }
3504  else if(expr.type().id() == ID_floatbv)
3505  {
3506  // Floating-point additions should have be been converted
3507  // to ID_floatbv_plus during symbolic execution, adding
3508  // the rounding mode. See smt2_convt::convert_floatbv_plus.
3509  UNREACHABLE;
3510  }
3511  else if(expr.type().id() == ID_pointer)
3512  {
3513  if(expr.operands().size() == 2)
3514  {
3515  exprt p = expr.op0(), i = expr.op1();
3516 
3517  if(p.type().id() != ID_pointer)
3518  p.swap(i);
3519 
3521  p.type().id() == ID_pointer,
3522  "one of the operands should have pointer type");
3523 
3524  const auto &base_type = to_pointer_type(expr.type()).base_type();
3525 
3526  mp_integer element_size;
3527  if(base_type.id() == ID_empty)
3528  {
3529  // This is a gcc extension.
3530  // https://gcc.gnu.org/onlinedocs/gcc-4.8.0/gcc/Pointer-Arith.html
3531  element_size = 1;
3532  }
3533  else
3534  {
3535  auto element_size_opt = pointer_offset_size(base_type, ns);
3536  CHECK_RETURN(element_size_opt.has_value() && *element_size_opt >= 1);
3537  element_size = *element_size_opt;
3538  }
3539 
3540  out << "(bvadd ";
3541  convert_expr(p);
3542  out << " ";
3543 
3544  if(element_size >= 2)
3545  {
3546  out << "(bvmul ";
3547  convert_expr(i);
3548  out << " (_ bv" << element_size << " " << boolbv_width(expr.type())
3549  << "))";
3550  }
3551  else
3552  convert_expr(i);
3553 
3554  out << ')';
3555  }
3556  else
3557  {
3559  }
3560  }
3561  else if(expr.type().id() == ID_vector)
3562  {
3563  const vector_typet &vector_type = to_vector_type(expr.type());
3564 
3565  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
3566 
3567  typet index_type = vector_type.size().type();
3568 
3569  if(use_datatypes)
3570  {
3571  const std::string &smt_typename = datatype_map.at(vector_type);
3572 
3573  out << "(mk-" << smt_typename;
3574  }
3575  else
3576  out << "(concat";
3577 
3578  // add component-by-component
3579  for(mp_integer i = 0; i != size; ++i)
3580  {
3581  exprt::operandst summands;
3582  summands.reserve(expr.operands().size());
3583  for(const auto &op : expr.operands())
3584  summands.push_back(index_exprt(
3585  op,
3586  from_integer(size - i - 1, index_type),
3587  vector_type.element_type()));
3588 
3589  plus_exprt tmp(std::move(summands), vector_type.element_type());
3590 
3591  out << " ";
3592  convert_expr(tmp);
3593  }
3594 
3595  out << ")"; // mk-... or concat
3596  }
3597  else
3598  UNEXPECTEDCASE("unsupported type for +: " + expr.type().id_string());
3599 }
3600 
3605 {
3607 
3608  /* CProver uses the x86 numbering of the rounding-mode
3609  * 0 == FE_TONEAREST
3610  * 1 == FE_DOWNWARD
3611  * 2 == FE_UPWARD
3612  * 3 == FE_TOWARDZERO
3613  * These literal values must be used rather than the macros
3614  * the macros from fenv.h to avoid portability problems.
3615  */
3616 
3617  if(expr.id()==ID_constant)
3618  {
3619  const constant_exprt &cexpr=to_constant_expr(expr);
3620 
3621  mp_integer value = numeric_cast_v<mp_integer>(cexpr);
3622 
3623  if(value==0)
3624  out << "roundNearestTiesToEven";
3625  else if(value==1)
3626  out << "roundTowardNegative";
3627  else if(value==2)
3628  out << "roundTowardPositive";
3629  else if(value==3)
3630  out << "roundTowardZero";
3631  else
3633  false,
3634  "Rounding mode should have value 0, 1, 2, or 3",
3635  id2string(cexpr.get_value()));
3636  }
3637  else
3638  {
3639  std::size_t width=to_bitvector_type(expr.type()).get_width();
3640 
3641  // Need to make the choice above part of the model
3642  out << "(ite (= (_ bv0 " << width << ") ";
3643  convert_expr(expr);
3644  out << ") roundNearestTiesToEven ";
3645 
3646  out << "(ite (= (_ bv1 " << width << ") ";
3647  convert_expr(expr);
3648  out << ") roundTowardNegative ";
3649 
3650  out << "(ite (= (_ bv2 " << width << ") ";
3651  convert_expr(expr);
3652  out << ") roundTowardPositive ";
3653 
3654  // TODO: add some kind of error checking here
3655  out << "roundTowardZero";
3656 
3657  out << ")))";
3658  }
3659 }
3660 
3662 {
3663  const typet &type=expr.type();
3664 
3665  PRECONDITION(
3666  type.id() == ID_floatbv ||
3667  (type.id() == ID_complex &&
3668  to_complex_type(type).subtype().id() == ID_floatbv) ||
3669  (type.id() == ID_vector &&
3670  to_vector_type(type).element_type().id() == ID_floatbv));
3671 
3672  if(use_FPA_theory)
3673  {
3674  if(type.id()==ID_floatbv)
3675  {
3676  out << "(fp.add ";
3678  out << " ";
3679  convert_expr(expr.lhs());
3680  out << " ";
3681  convert_expr(expr.rhs());
3682  out << ")";
3683  }
3684  else if(type.id()==ID_complex)
3685  {
3686  SMT2_TODO("+ for floatbv complex");
3687  }
3688  else if(type.id()==ID_vector)
3689  {
3690  SMT2_TODO("+ for floatbv vector");
3691  }
3692  else
3694  false,
3695  "type should not be one of the unsupported types",
3696  type.id_string());
3697  }
3698  else
3699  convert_floatbv(expr);
3700 }
3701 
3703 {
3704  if(expr.type().id()==ID_integer)
3705  {
3706  out << "(- ";
3707  convert_expr(expr.op0());
3708  out << " ";
3709  convert_expr(expr.op1());
3710  out << ")";
3711  }
3712  else if(expr.type().id()==ID_unsignedbv ||
3713  expr.type().id()==ID_signedbv ||
3714  expr.type().id()==ID_fixedbv)
3715  {
3716  if(expr.op0().type().id()==ID_pointer &&
3717  expr.op1().type().id()==ID_pointer)
3718  {
3719  // Pointer difference
3720  auto element_size =
3722  CHECK_RETURN(element_size.has_value() && *element_size >= 1);
3723 
3724  if(*element_size >= 2)
3725  out << "(bvsdiv ";
3726 
3727  INVARIANT(
3728  boolbv_width(expr.op0().type()) == boolbv_width(expr.type()),
3729  "bitvector width of operand shall be equal to the bitvector width of "
3730  "the expression");
3731 
3732  out << "(bvsub ";
3733  convert_expr(expr.op0());
3734  out << " ";
3735  convert_expr(expr.op1());
3736  out << ")";
3737 
3738  if(*element_size >= 2)
3739  out << " (_ bv" << *element_size << " " << boolbv_width(expr.type())
3740  << "))";
3741  }
3742  else
3743  {
3744  out << "(bvsub ";
3745  convert_expr(expr.op0());
3746  out << " ";
3747  convert_expr(expr.op1());
3748  out << ")";
3749  }
3750  }
3751  else if(expr.type().id()==ID_floatbv)
3752  {
3753  // Floating-point subtraction should have be been converted
3754  // to ID_floatbv_minus during symbolic execution, adding
3755  // the rounding mode. See smt2_convt::convert_floatbv_minus.
3756  UNREACHABLE;
3757  }
3758  else if(expr.type().id()==ID_pointer)
3759  {
3760  if(
3761  expr.op0().type().id() == ID_pointer &&
3762  (expr.op1().type().id() == ID_unsignedbv ||
3763  expr.op1().type().id() == ID_signedbv))
3764  {
3765  // rewrite p-o to p+(-o)
3766  return convert_plus(
3767  plus_exprt(expr.op0(), unary_minus_exprt(expr.op1())));
3768  }
3769  else
3771  "unsupported operand types for -: " + expr.op0().type().id_string() +
3772  " and " + expr.op1().type().id_string());
3773  }
3774  else if(expr.type().id()==ID_vector)
3775  {
3776  const vector_typet &vector_type=to_vector_type(expr.type());
3777 
3778  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
3779 
3780  typet index_type=vector_type.size().type();
3781 
3782  if(use_datatypes)
3783  {
3784  const std::string &smt_typename = datatype_map.at(vector_type);
3785 
3786  out << "(mk-" << smt_typename;
3787  }
3788  else
3789  out << "(concat";
3790 
3791  // subtract component-by-component
3792  for(mp_integer i=0; i!=size; ++i)
3793  {
3794  exprt tmp(ID_minus, vector_type.element_type());
3795  forall_operands(it, expr)
3797  *it,
3798  from_integer(size - i - 1, index_type),
3799  vector_type.element_type()));
3800 
3801  out << " ";
3802  convert_expr(tmp);
3803  }
3804 
3805  out << ")"; // mk-... or concat
3806  }
3807  else
3808  UNEXPECTEDCASE("unsupported type for -: "+expr.type().id_string());
3809 }
3810 
3812 {
3814  expr.type().id() == ID_floatbv,
3815  "type of ieee floating point expression shall be floatbv");
3816 
3817  if(use_FPA_theory)
3818  {
3819  out << "(fp.sub ";
3821  out << " ";
3822  convert_expr(expr.lhs());
3823  out << " ";
3824  convert_expr(expr.rhs());
3825  out << ")";
3826  }
3827  else
3828  convert_floatbv(expr);
3829 }
3830 
3832 {
3833  if(expr.type().id()==ID_unsignedbv ||
3834  expr.type().id()==ID_signedbv)
3835  {
3836  if(expr.type().id()==ID_unsignedbv)
3837  out << "(bvudiv ";
3838  else
3839  out << "(bvsdiv ";
3840 
3841  convert_expr(expr.op0());
3842  out << " ";
3843  convert_expr(expr.op1());
3844  out << ")";
3845  }
3846  else if(expr.type().id()==ID_fixedbv)
3847  {
3848  fixedbv_spect spec(to_fixedbv_type(expr.type()));
3849  std::size_t fraction_bits=spec.get_fraction_bits();
3850 
3851  out << "((_ extract " << spec.width-1 << " 0) ";
3852  out << "(bvsdiv ";
3853 
3854  out << "(concat ";
3855  convert_expr(expr.op0());
3856  out << " (_ bv0 " << fraction_bits << ")) ";
3857 
3858  out << "((_ sign_extend " << fraction_bits << ") ";
3859  convert_expr(expr.op1());
3860  out << ")";
3861 
3862  out << "))";
3863  }
3864  else if(expr.type().id()==ID_floatbv)
3865  {
3866  // Floating-point division should have be been converted
3867  // to ID_floatbv_div during symbolic execution, adding
3868  // the rounding mode. See smt2_convt::convert_floatbv_div.
3869  UNREACHABLE;
3870  }
3871  else
3872  UNEXPECTEDCASE("unsupported type for /: "+expr.type().id_string());
3873 }
3874 
3876 {
3878  expr.type().id() == ID_floatbv,
3879  "type of ieee floating point expression shall be floatbv");
3880 
3881  if(use_FPA_theory)
3882  {
3883  out << "(fp.div ";
3885  out << " ";
3886  convert_expr(expr.lhs());
3887  out << " ";
3888  convert_expr(expr.rhs());
3889  out << ")";
3890  }
3891  else
3892  convert_floatbv(expr);
3893 }
3894 
3896 {
3897  // re-write to binary if needed
3898  if(expr.operands().size()>2)
3899  {
3900  // strip last operand
3901  exprt tmp=expr;
3902  tmp.operands().pop_back();
3903 
3904  // recursive call
3905  return convert_mult(mult_exprt(tmp, expr.operands().back()));
3906  }
3907 
3908  INVARIANT(
3909  expr.operands().size() == 2,
3910  "expression should have been converted to a variant with two operands");
3911 
3912  if(expr.type().id()==ID_unsignedbv ||
3913  expr.type().id()==ID_signedbv)
3914  {
3915  // Note that bvmul is really unsigned,
3916  // but this is irrelevant as we chop-off any extra result
3917  // bits.
3918  out << "(bvmul ";
3919  convert_expr(expr.op0());
3920  out << " ";
3921  convert_expr(expr.op1());
3922  out << ")";
3923  }
3924  else if(expr.type().id()==ID_floatbv)
3925  {
3926  // Floating-point multiplication should have be been converted
3927  // to ID_floatbv_mult during symbolic execution, adding
3928  // the rounding mode. See smt2_convt::convert_floatbv_mult.
3929  UNREACHABLE;
3930  }
3931  else if(expr.type().id()==ID_fixedbv)
3932  {
3933  fixedbv_spect spec(to_fixedbv_type(expr.type()));
3934  std::size_t fraction_bits=spec.get_fraction_bits();
3935 
3936  out << "((_ extract "
3937  << spec.width+fraction_bits-1 << " "
3938  << fraction_bits << ") ";
3939 
3940  out << "(bvmul ";
3941 
3942  out << "((_ sign_extend " << fraction_bits << ") ";
3943  convert_expr(expr.op0());
3944  out << ") ";
3945 
3946  out << "((_ sign_extend " << fraction_bits << ") ";
3947  convert_expr(expr.op1());
3948  out << ")";
3949 
3950  out << "))"; // bvmul, extract
3951  }
3952  else if(expr.type().id()==ID_rational ||
3953  expr.type().id()==ID_integer ||
3954  expr.type().id()==ID_real)
3955  {
3956  out << "(*";
3957 
3958  forall_operands(it, expr)
3959  {
3960  out << " ";
3961  convert_expr(*it);
3962  }
3963 
3964  out << ")";
3965  }
3966  else
3967  UNEXPECTEDCASE("unsupported type for *: "+expr.type().id_string());
3968 }
3969 
3971 {
3973  expr.type().id() == ID_floatbv,
3974  "type of ieee floating point expression shall be floatbv");
3975 
3976  if(use_FPA_theory)
3977  {
3978  out << "(fp.mul ";
3980  out << " ";
3981  convert_expr(expr.lhs());
3982  out << " ";
3983  convert_expr(expr.rhs());
3984  out << ")";
3985  }
3986  else
3987  convert_floatbv(expr);
3988 }
3989 
3991 {
3993  expr.type().id() == ID_floatbv,
3994  "type of ieee floating point expression shall be floatbv");
3995 
3996  if(use_FPA_theory)
3997  {
3998  // Note that these do not have a rounding mode
3999  out << "(fp.rem ";
4000  convert_expr(expr.lhs());
4001  out << " ";
4002  convert_expr(expr.rhs());
4003  out << ")";
4004  }
4005  else
4006  {
4007  SMT2_TODO(
4008  "smt2_convt::convert_floatbv_rem to be implemented when not using "
4009  "FPA_theory");
4010  }
4011 }
4012 
4014 {
4015  // get rid of "with" that has more than three operands
4016 
4017  if(expr.operands().size()>3)
4018  {
4019  std::size_t s=expr.operands().size();
4020 
4021  // strip off the trailing two operands
4022  with_exprt tmp = expr;
4023  tmp.operands().resize(s-2);
4024 
4025  with_exprt new_with_expr(
4026  tmp, expr.operands()[s - 2], expr.operands().back());
4027 
4028  // recursive call
4029  return convert_with(new_with_expr);
4030  }
4031 
4032  INVARIANT(
4033  expr.operands().size() == 3,
4034  "with expression should have been converted to a version with three "
4035  "operands above");
4036 
4037  const typet &expr_type = expr.type();
4038 
4039  if(expr_type.id()==ID_array)
4040  {
4041  const array_typet &array_type=to_array_type(expr_type);
4042 
4043  if(use_array_theory(expr))
4044  {
4045  out << "(store ";
4046  convert_expr(expr.old());
4047  out << " ";
4048  convert_expr(typecast_exprt(expr.where(), array_type.index_type()));
4049  out << " ";
4050  convert_expr(expr.new_value());
4051  out << ")";
4052  }
4053  else
4054  {
4055  // fixed-width
4056  std::size_t array_width=boolbv_width(array_type);
4057  std::size_t sub_width = boolbv_width(array_type.element_type());
4058  std::size_t index_width=boolbv_width(expr.where().type());
4059 
4060  // We mask out the updated bits with AND,
4061  // and then OR-in the shifted new value.
4062 
4063  out << "(let ((distance? ";
4064  out << "(bvmul (_ bv" << sub_width << " " << array_width << ") ";
4065 
4066  // SMT2 says that the shift distance needs to be as wide
4067  // as the stuff we are shifting.
4068  if(array_width>index_width)
4069  {
4070  out << "((_ zero_extend " << array_width-index_width << ") ";
4071  convert_expr(expr.where());
4072  out << ")";
4073  }
4074  else
4075  {
4076  out << "((_ extract " << array_width-1 << " 0) ";
4077  convert_expr(expr.where());
4078  out << ")";
4079  }
4080 
4081  out << "))) "; // bvmul, distance?
4082 
4083  out << "(bvor ";
4084  out << "(bvand ";
4085  out << "(bvnot ";
4086  out << "(bvshl (_ bv" << power(2, sub_width) - 1 << " " << array_width
4087  << ") ";
4088  out << "distance?)) "; // bvnot, bvlshl
4089  convert_expr(expr.old());
4090  out << ") "; // bvand
4091  out << "(bvshl ";
4092  out << "((_ zero_extend " << array_width-sub_width << ") ";
4093  convert_expr(expr.new_value());
4094  out << ") distance?)))"; // zero_extend, bvshl, bvor, let
4095  }
4096  }
4097  else if(expr_type.id() == ID_struct || expr_type.id() == ID_struct_tag)
4098  {
4099  const struct_typet &struct_type = to_struct_type(ns.follow(expr_type));
4100 
4101  const exprt &index = expr.where();
4102  const exprt &value = expr.new_value();
4103 
4104  const irep_idt &component_name=index.get(ID_component_name);
4105 
4106  INVARIANT(
4107  struct_type.has_component(component_name),
4108  "struct should have accessed component");
4109 
4110  if(use_datatypes)
4111  {
4112  const std::string &smt_typename = datatype_map.at(expr_type);
4113 
4114  out << "(update-" << smt_typename << "." << component_name << " ";
4115  convert_expr(expr.old());
4116  out << " ";
4117  convert_expr(value);
4118  out << ")";
4119  }
4120  else
4121  {
4122  std::size_t struct_width=boolbv_width(struct_type);
4123 
4124  // figure out the offset and width of the member
4125  const boolbv_widtht::membert &m =
4126  boolbv_width.get_member(struct_type, component_name);
4127 
4128  out << "(let ((?withop ";
4129  convert_expr(expr.old());
4130  out << ")) ";
4131 
4132  if(m.width==struct_width)
4133  {
4134  // the struct is the same as the member, no concat needed,
4135  // ?withop won't be used
4136  convert_expr(value);
4137  }
4138  else if(m.offset==0)
4139  {
4140  // the member is at the beginning
4141  out << "(concat "
4142  << "((_ extract " << (struct_width-1) << " "
4143  << m.width << ") ?withop) ";
4144  convert_expr(value);
4145  out << ")"; // concat
4146  }
4147  else if(m.offset+m.width==struct_width)
4148  {
4149  // the member is at the end
4150  out << "(concat ";
4151  convert_expr(value);
4152  out << " ((_ extract " << (m.offset - 1) << " 0) ?withop))";
4153  }
4154  else
4155  {
4156  // most general case, need two concat-s
4157  out << "(concat (concat "
4158  << "((_ extract " << (struct_width-1) << " "
4159  << (m.offset+m.width) << ") ?withop) ";
4160  convert_expr(value);
4161  out << ") ((_ extract " << (m.offset-1) << " 0) ?withop)";
4162  out << ")"; // concat
4163  }
4164 
4165  out << ")"; // let ?withop
4166  }
4167  }
4168  else if(expr_type.id() == ID_union || expr_type.id() == ID_union_tag)
4169  {
4170  const union_typet &union_type = to_union_type(ns.follow(expr_type));
4171 
4172  const exprt &value = expr.new_value();
4173 
4174  std::size_t total_width=boolbv_width(union_type);
4175 
4176  std::size_t member_width=boolbv_width(value.type());
4177 
4178  if(total_width==member_width)
4179  {
4180  flatten2bv(value);
4181  }
4182  else
4183  {
4184  INVARIANT(
4185  total_width > member_width,
4186  "total width should be greater than member_width as member_width is at "
4187  "most as large as total_width and the other case has been handled "
4188  "above");
4189  out << "(concat ";
4190  out << "((_ extract "
4191  << (total_width-1)
4192  << " " << member_width << ") ";
4193  convert_expr(expr.old());
4194  out << ") ";
4195  flatten2bv(value);
4196  out << ")";
4197  }
4198  }
4199  else if(expr_type.id()==ID_bv ||
4200  expr_type.id()==ID_unsignedbv ||
4201  expr_type.id()==ID_signedbv)
4202  {
4203  // Update bits in a bit-vector. We will use masking and shifts.
4204  // TODO: SMT2-ify
4205  SMT2_TODO("SMT2-ify");
4206 
4207 #if 0
4208  std::size_t total_width=boolbv_width(expr_type);
4209 
4210  const exprt &index=expr.operands()[1];
4211  const exprt &value=expr.operands()[2];
4212 
4213  std::size_t value_width=boolbv_width(value.type());
4214 
4215  typecast_exprt index_tc(index, expr_type);
4216 
4217  out << "(bvor ";
4218  out << "(band ";
4219 
4220  // the mask to get rid of the old bits
4221  out << " (bvnot (bvshl";
4222 
4223  out << " (concat";
4224  out << " (repeat[" << total_width-value_width << "] bv0[1])";
4225  out << " (repeat[" << value_width << "] bv1[1])";
4226  out << ")"; // concat
4227 
4228  // shift it to the index
4229  convert_expr(index_tc);
4230  out << "))"; // bvshl, bvot
4231 
4232  out << ")"; // bvand
4233 
4234  // the new value
4235  out << " (bvshl ";
4236  convert_expr(value);
4237 
4238  // shift it to the index
4239  convert_expr(index_tc);
4240  out << ")"; // bvshl
4241 
4242  out << ")"; // bvor
4243 #endif
4244  }
4245  else
4247  "with expects struct, union, or array type, but got "+
4248  expr.type().id_string());
4249 }
4250 
4252 {
4253  PRECONDITION(expr.operands().size() == 3);
4254 
4255  SMT2_TODO("smt2_convt::convert_update to be implemented");
4256 }
4257 
4259 {
4260  const typet &array_op_type = expr.array().type();
4261 
4262  if(array_op_type.id()==ID_array)
4263  {
4264  const array_typet &array_type=to_array_type(array_op_type);
4265 
4266  if(use_array_theory(expr.array()))
4267  {
4268  if(expr.type().id() == ID_bool && !use_array_of_bool)
4269  {
4270  out << "(= ";
4271  out << "(select ";
4272  convert_expr(expr.array());
4273  out << " ";
4274  convert_expr(typecast_exprt(expr.index(), array_type.index_type()));
4275  out << ")";
4276  out << " #b1)";
4277  }
4278  else
4279  {
4280  out << "(select ";
4281  convert_expr(expr.array());
4282  out << " ";
4283  convert_expr(typecast_exprt(expr.index(), array_type.index_type()));
4284  out << ")";
4285  }
4286  }
4287  else
4288  {
4289  // fixed size
4290  std::size_t array_width=boolbv_width(array_type);
4291 
4292  unflatten(wheret::BEGIN, array_type.element_type());
4293 
4294  std::size_t sub_width = boolbv_width(array_type.element_type());
4295  std::size_t index_width=boolbv_width(expr.index().type());
4296 
4297  out << "((_ extract " << sub_width-1 << " 0) ";
4298  out << "(bvlshr ";
4299  convert_expr(expr.array());
4300  out << " ";
4301  out << "(bvmul (_ bv" << sub_width << " " << array_width << ") ";
4302 
4303  // SMT2 says that the shift distance must be the same as
4304  // the width of what we shift.
4305  if(array_width>index_width)
4306  {
4307  out << "((_ zero_extend " << array_width-index_width << ") ";
4308  convert_expr(expr.index());
4309  out << ")"; // zero_extend
4310  }
4311  else
4312  {
4313  out << "((_ extract " << array_width-1 << " 0) ";
4314  convert_expr(expr.index());
4315  out << ")"; // extract
4316  }
4317 
4318  out << ")))"; // mult, bvlshr, extract
4319 
4320  unflatten(wheret::END, array_type.element_type());
4321  }
4322  }
4323  else if(array_op_type.id()==ID_vector)
4324  {
4325  const vector_typet &vector_type=to_vector_type(array_op_type);
4326 
4327  if(use_datatypes)
4328  {
4329  const std::string &smt_typename = datatype_map.at(vector_type);
4330 
4331  // this is easy for constant indicies
4332 
4333  const auto index_int = numeric_cast<mp_integer>(expr.index());
4334  if(!index_int.has_value())
4335  {
4336  SMT2_TODO("non-constant index on vectors");
4337  }
4338  else
4339  {
4340  out << "(" << smt_typename << "." << *index_int << " ";
4341  convert_expr(expr.array());
4342  out << ")";
4343  }
4344  }
4345  else
4346  {
4347  SMT2_TODO("index on vectors");
4348  }
4349  }
4350  else
4351  INVARIANT(
4352  false, "index with unsupported array type: " + array_op_type.id_string());
4353 }
4354 
4356 {
4357  const member_exprt &member_expr=to_member_expr(expr);
4358  const exprt &struct_op=member_expr.struct_op();
4359  const typet &struct_op_type = struct_op.type();
4360  const irep_idt &name=member_expr.get_component_name();
4361 
4362  if(struct_op_type.id() == ID_struct || struct_op_type.id() == ID_struct_tag)
4363  {
4364  const struct_typet &struct_type = to_struct_type(ns.follow(struct_op_type));
4365 
4366  INVARIANT(
4367  struct_type.has_component(name), "struct should have accessed component");
4368 
4369  if(use_datatypes)
4370  {
4371  const std::string &smt_typename = datatype_map.at(struct_type);
4372 
4373  out << "(" << smt_typename << "."
4374  << struct_type.get_component(name).get_name()
4375  << " ";
4376  convert_expr(struct_op);
4377  out << ")";
4378  }
4379  else
4380  {
4381  // we extract
4382  const auto &member_offset = boolbv_width.get_member(struct_type, name);
4383 
4384  if(expr.type().id() == ID_bool)
4385  out << "(= ";
4386  out << "((_ extract " << (member_offset.offset + member_offset.width - 1)
4387  << " " << member_offset.offset << ") ";
4388  convert_expr(struct_op);
4389  out << ")";
4390  if(expr.type().id() == ID_bool)
4391  out << " #b1)";
4392  }
4393  }
4394  else if(
4395  struct_op_type.id() == ID_union || struct_op_type.id() == ID_union_tag)
4396  {
4397  std::size_t width=boolbv_width(expr.type());
4399  width != 0, "failed to get union member width");
4400 
4401  unflatten(wheret::BEGIN, expr.type());
4402 
4403  out << "((_ extract "
4404  << (width-1)
4405  << " 0) ";
4406  convert_expr(struct_op);
4407  out << ")";
4408 
4409  unflatten(wheret::END, expr.type());
4410  }
4411  else
4413  "convert_member on an unexpected type "+struct_op_type.id_string());
4414 }
4415 
4417 {
4418  const typet &type = expr.type();
4419 
4420  if(type.id()==ID_bool)
4421  {
4422  out << "(ite ";
4423  convert_expr(expr); // this returns a Bool
4424  out << " #b1 #b0)"; // this is a one-bit bit-vector
4425  }
4426  else if(type.id()==ID_vector)
4427  {
4428  if(use_datatypes)
4429  {
4430  const std::string &smt_typename = datatype_map.at(type);
4431 
4432  // concatenate elements
4433  const vector_typet &vector_type=to_vector_type(type);
4434 
4435  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
4436 
4437  out << "(let ((?vflop ";
4438  convert_expr(expr);
4439  out << ")) ";
4440 
4441  out << "(concat";
4442 
4443  for(mp_integer i=0; i!=size; ++i)
4444  {
4445  out << " (" << smt_typename << "." << i << " ?vflop)";
4446  }
4447 
4448  out << "))"; // concat, let
4449  }
4450  else
4451  convert_expr(expr); // already a bv
4452  }
4453  else if(type.id()==ID_array)
4454  {
4455  convert_expr(expr);
4456  }
4457  else if(type.id() == ID_struct || type.id() == ID_struct_tag)
4458  {
4459  if(use_datatypes)
4460  {
4461  const std::string &smt_typename = datatype_map.at(type);
4462 
4463  // concatenate elements
4464  const struct_typet &struct_type = to_struct_type(ns.follow(type));
4465 
4466  out << "(let ((?sflop ";
4467  convert_expr(expr);
4468  out << ")) ";
4469 
4470  const struct_typet::componentst &components=
4471  struct_type.components();
4472 
4473  // SMT-LIB 2 concat is binary only
4474  for(std::size_t i=components.size(); i>1; i--)
4475  {
4476  out << "(concat (" << smt_typename << "."
4477  << components[i-1].get_name() << " ?sflop)";
4478 
4479  out << " ";
4480  }
4481 
4482  out << "(" << smt_typename << "."
4483  << components[0].get_name() << " ?sflop)";
4484 
4485  for(std::size_t i=1; i<components.size(); i++)
4486  out << ")"; // concat
4487 
4488  out << ")"; // let
4489  }
4490  else
4491  convert_expr(expr);
4492  }
4493  else if(type.id()==ID_floatbv)
4494  {
4495  INVARIANT(
4496  !use_FPA_theory,
4497  "floatbv expressions should be flattened when using FPA theory");
4498 
4499  convert_expr(expr);
4500  }
4501  else
4502  convert_expr(expr);
4503 }
4504 
4506  wheret where,
4507  const typet &type,
4508  unsigned nesting)
4509 {
4510  if(type.id()==ID_bool)
4511  {
4512  if(where==wheret::BEGIN)
4513  out << "(= "; // produces a bool
4514  else
4515  out << " #b1)";
4516  }
4517  else if(type.id()==ID_vector)
4518  {
4519  if(use_datatypes)
4520  {
4521  const std::string &smt_typename = datatype_map.at(type);
4522 
4523  // extract elements
4524  const vector_typet &vector_type=to_vector_type(type);
4525 
4526  std::size_t subtype_width = boolbv_width(vector_type.element_type());
4527 
4528  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
4529 
4530  if(where==wheret::BEGIN)
4531  out << "(let ((?ufop" << nesting << " ";
4532  else
4533  {
4534  out << ")) ";
4535 
4536  out << "(mk-" << smt_typename;
4537 
4538  std::size_t offset=0;
4539 
4540  for(mp_integer i=0; i!=size; ++i, offset+=subtype_width)
4541  {
4542  out << " ";
4543  unflatten(wheret::BEGIN, vector_type.element_type(), nesting + 1);
4544  out << "((_ extract " << offset+subtype_width-1 << " "
4545  << offset << ") ?ufop" << nesting << ")";
4546  unflatten(wheret::END, vector_type.element_type(), nesting + 1);
4547  }
4548 
4549  out << "))"; // mk-, let
4550  }
4551  }
4552  else
4553  {
4554  // nop, already a bv
4555  }
4556  }
4557  else if(type.id() == ID_struct || type.id() == ID_struct_tag)
4558  {
4559  if(use_datatypes)
4560  {
4561  // extract members
4562  if(where==wheret::BEGIN)
4563  out << "(let ((?ufop" << nesting << " ";
4564  else
4565  {
4566  out << ")) ";
4567 
4568  const std::string &smt_typename = datatype_map.at(type);
4569 
4570  out << "(mk-" << smt_typename;
4571 
4572  const struct_typet &struct_type = to_struct_type(ns.follow(type));
4573 
4574  const struct_typet::componentst &components=
4575  struct_type.components();
4576 
4577  std::size_t offset=0;
4578 
4579  std::size_t i=0;
4580  for(struct_typet::componentst::const_iterator
4581  it=components.begin();
4582  it!=components.end();
4583  it++, i++)
4584  {
4585  std::size_t member_width=boolbv_width(it->type());
4586 
4587  out << " ";
4588  unflatten(wheret::BEGIN, it->type(), nesting+1);
4589  out << "((_ extract " << offset+member_width-1 << " "
4590  << offset << ") ?ufop" << nesting << ")";
4591  unflatten(wheret::END, it->type(), nesting+1);
4592  offset+=member_width;
4593  }
4594 
4595  out << "))"; // mk-, let
4596  }
4597  }
4598  else
4599  {
4600  // nop, already a bv
4601  }
4602  }
4603  else
4604  {
4605  // nop
4606  }
4607 }
4608 
4609 void smt2_convt::set_to(const exprt &expr, bool value)
4610 {
4611  PRECONDITION(expr.type().id() == ID_bool);
4612 
4613  if(expr.id()==ID_and && value)
4614  {
4615  forall_operands(it, expr)
4616  set_to(*it, true);
4617  return;
4618  }
4619 
4620  if(expr.id()==ID_or && !value)
4621  {
4622  forall_operands(it, expr)
4623  set_to(*it, false);
4624  return;
4625  }
4626 
4627  if(expr.id()==ID_not)
4628  {
4629  return set_to(to_not_expr(expr).op(), !value);
4630  }
4631 
4632  out << "\n";
4633 
4634  // special treatment for "set_to(a=b, true)" where
4635  // a is a new symbol
4636 
4637  if(expr.id() == ID_equal && value)
4638  {
4639  const equal_exprt &equal_expr=to_equal_expr(expr);
4640  if(
4641  equal_expr.lhs().type().id() == ID_empty ||
4642  equal_expr.rhs().id() == ID_empty_union)
4643  {
4644  // ignore equality checking over expressions with empty (void) type
4645  return;
4646  }
4647 
4648  if(equal_expr.lhs().id()==ID_symbol)
4649  {
4650  const irep_idt &identifier=
4651  to_symbol_expr(equal_expr.lhs()).get_identifier();
4652 
4653  if(identifier_map.find(identifier)==identifier_map.end())
4654  {
4655  identifiert &id=identifier_map[identifier];
4656  CHECK_RETURN(id.type.is_nil());
4657 
4658  id.type=equal_expr.lhs().type();
4659  find_symbols(id.type);
4660  exprt prepared_rhs = prepare_for_convert_expr(equal_expr.rhs());
4661 
4662  std::string smt2_identifier=convert_identifier(identifier);
4663  smt2_identifiers.insert(smt2_identifier);
4664 
4665  out << "; set_to true (equal)\n";
4666  out << "(define-fun " << smt2_identifier;
4667 
4668  if(equal_expr.lhs().type().id() == ID_mathematical_function)
4669  {
4670  auto &mathematical_function_type =
4671  to_mathematical_function_type(equal_expr.lhs().type());
4672  out << " (";
4673  bool first = true;
4674 
4675  for(std::size_t p_nr = 0;
4676  p_nr < mathematical_function_type.domain().size();
4677  p_nr++)
4678  {
4679  if(first)
4680  first = false;
4681  else
4682  out << ' ';
4683 
4684  out << '(' << 'p' << (p_nr + 1) << ' ';
4685  convert_type(mathematical_function_type.domain()[p_nr]);
4686  out << ')';
4687  }
4688 
4689  out << ") ";
4690  convert_type(mathematical_function_type.codomain());
4691 
4692  out << ' ' << '(';
4693  convert_expr(prepared_rhs);
4694  for(std::size_t p_nr = 0;
4695  p_nr < mathematical_function_type.domain().size();
4696  p_nr++)
4697  out << ' ' << 'p' << (p_nr + 1);
4698  out << ')';
4699  }
4700  else
4701  {
4702  out << " () ";
4703  convert_type(equal_expr.lhs().type());
4704  out << ' ';
4705  convert_expr(prepared_rhs);
4706  }
4707 
4708  out << ')' << '\n';
4709  return; // done
4710  }
4711  }
4712  }
4713 
4714  exprt prepared_expr = prepare_for_convert_expr(expr);
4715 
4716 #if 0
4717  out << "; CONV: "
4718  << format(expr) << "\n";
4719 #endif
4720 
4721  out << "; set_to " << (value?"true":"false") << "\n"
4722  << "(assert ";
4723  if(!value)
4724  {
4725  out << "(not ";
4726  }
4727  const auto found_literal = defined_expressions.find(expr);
4728  if(!(found_literal == defined_expressions.end()))
4729  {
4730  // This is a converted expression, we can just assert the literal name
4731  // since the expression is already defined
4732  out << found_literal->second;
4733  set_values[found_literal->second] = value;
4734  }
4735  else
4736  {
4737  convert_expr(prepared_expr);
4738  }
4739  if(!value)
4740  {
4741  out << ")";
4742  }
4743  out << ")\n";
4744  return;
4745 }
4746 
4754 {
4755  exprt lowered_expr = expr;
4756 
4757  for(auto it = lowered_expr.depth_begin(), itend = lowered_expr.depth_end();
4758  it != itend;
4759  ++it)
4760  {
4761  if(
4762  it->id() == ID_byte_extract_little_endian ||
4763  it->id() == ID_byte_extract_big_endian)
4764  {
4765  it.mutate() = lower_byte_extract(to_byte_extract_expr(*it), ns);
4766  }
4767  else if(
4768  it->id() == ID_byte_update_little_endian ||
4769  it->id() == ID_byte_update_big_endian)
4770  {
4771  it.mutate() = lower_byte_update(to_byte_update_expr(*it), ns);
4772  }
4773  }
4774 
4775  return lowered_expr;
4776 }
4777 
4786 {
4787  // First, replace byte operators, because they may introduce new array
4788  // expressions that must be seen by find_symbols:
4789  exprt lowered_expr = lower_byte_operators(expr);
4790  INVARIANT(
4791  !has_byte_operator(lowered_expr),
4792  "lower_byte_operators should remove all byte operators");
4793 
4794  // Now create symbols for all composite expressions present in lowered_expr:
4795  find_symbols(lowered_expr);
4796 
4797  return lowered_expr;
4798 }
4799 
4801 {
4802  // recursive call on type
4803  find_symbols(expr.type());
4804 
4805  if(expr.id() == ID_exists || expr.id() == ID_forall)
4806  {
4807  // do not declare the quantified symbol, but record
4808  // as 'bound symbol'
4809  const auto &q_expr = to_quantifier_expr(expr);
4810  for(const auto &symbol : q_expr.variables())
4811  {
4812  const auto identifier = symbol.get_identifier();
4813  identifiert &id = identifier_map[identifier];
4814  id.type = symbol.type();
4815  id.is_bound = true;
4816  }
4817  find_symbols(q_expr.where());
4818  return;
4819  }
4820 
4821  // recursive call on operands
4822  forall_operands(it, expr)
4823  find_symbols(*it);
4824 
4825  if(expr.id()==ID_symbol ||
4826  expr.id()==ID_nondet_symbol)
4827  {
4828  // we don't track function-typed symbols
4829  if(expr.type().id()==ID_code)
4830  return;
4831 
4832  irep_idt identifier;
4833 
4834  if(expr.id()==ID_symbol)
4835  identifier=to_symbol_expr(expr).get_identifier();
4836  else
4837  identifier="nondet_"+
4839 
4840  identifiert &id=identifier_map[identifier];
4841 
4842  if(id.type.is_nil())
4843  {
4844  id.type=expr.type();
4845 
4846  std::string smt2_identifier=convert_identifier(identifier);
4847  smt2_identifiers.insert(smt2_identifier);
4848 
4849  out << "; find_symbols\n";
4850  out << "(declare-fun " << smt2_identifier;
4851 
4852  if(expr.type().id() == ID_mathematical_function)
4853  {
4854  auto &mathematical_function_type =
4856  out << " (";
4857  bool first = true;
4858 
4859  for(auto &type : mathematical_function_type.domain())
4860  {
4861  if(first)
4862  first = false;
4863  else
4864  out << ' ';
4865  convert_type(type);
4866  }
4867 
4868  out << ") ";
4869  convert_type(mathematical_function_type.codomain());
4870  }
4871  else
4872  {
4873  out << " () ";
4874  convert_type(expr.type());
4875  }
4876 
4877  out << ")" << "\n";
4878  }
4879  }
4880  else if(expr.id() == ID_array_of)
4881  {
4882  if(!use_as_const)
4883  {
4884  if(defined_expressions.find(expr) == defined_expressions.end())
4885  {
4886  const auto &array_of = to_array_of_expr(expr);
4887  const auto &array_type = array_of.type();
4888 
4889  const irep_idt id =
4890  "array_of." + std::to_string(defined_expressions.size());
4891  out << "; the following is a substitute for lambda i. x\n";
4892  out << "(declare-fun " << id << " () ";
4893  convert_type(array_type);
4894  out << ")\n";
4895 
4896  // use a quantifier-based initialization instead of lambda
4897  out << "(assert (forall ((i ";
4898  convert_type(array_type.index_type());
4899  out << ")) (= (select " << id << " i) ";
4900  if(array_type.element_type().id() == ID_bool && !use_array_of_bool)
4901  {
4902  out << "(ite ";
4903  convert_expr(array_of.what());
4904  out << " #b1 #b0)";
4905  }
4906  else
4907  {
4908  convert_expr(array_of.what());
4909  }
4910  out << ")))\n";
4911 
4912  defined_expressions[expr] = id;
4913  }
4914  }
4915  }
4916  else if(expr.id() == ID_array_comprehension)
4917  {
4919  {
4920  if(defined_expressions.find(expr) == defined_expressions.end())
4921  {
4922  const auto &array_comprehension = to_array_comprehension_expr(expr);
4923  const auto &array_type = array_comprehension.type();
4924  const auto &array_size = array_type.size();
4925 
4926  const irep_idt id =
4927  "array_comprehension." + std::to_string(defined_expressions.size());
4928  out << "(declare-fun " << id << " () ";
4929  convert_type(array_type);
4930  out << ")\n";
4931 
4932  out << "; the following is a substitute for lambda i . x(i)\n";
4933  out << "; universally quantified initialization of the array\n";
4934  out << "(assert (forall ((";
4935  convert_expr(array_comprehension.arg());
4936  out << " ";
4937  convert_type(array_size.type());
4938  out << ")) (=> (and (bvule (_ bv0 " << boolbv_width(array_size.type())
4939  << ") ";
4940  convert_expr(array_comprehension.arg());
4941  out << ") (bvult ";
4942  convert_expr(array_comprehension.arg());
4943  out << " ";
4944  convert_expr(array_size);
4945  out << ")) (= (select " << id << " ";
4946  convert_expr(array_comprehension.arg());
4947  out << ") ";
4948  if(array_type.element_type().id() == ID_bool && !use_array_of_bool)
4949  {
4950  out << "(ite ";
4951  convert_expr(array_comprehension.body());
4952  out << " #b1 #b0)";
4953  }
4954  else
4955  {
4956  convert_expr(array_comprehension.body());
4957  }
4958  out << "))))\n";
4959 
4960  defined_expressions[expr] = id;
4961  }
4962  }
4963  }
4964  else if(expr.id()==ID_array)
4965  {
4966  if(defined_expressions.find(expr)==defined_expressions.end())
4967  {
4968  const array_typet &array_type=to_array_type(expr.type());
4969 
4970  const irep_idt id = "array." + std::to_string(defined_expressions.size());
4971  out << "; the following is a substitute for an array constructor" << "\n";
4972  out << "(declare-fun " << id << " () ";
4973  convert_type(array_type);
4974  out << ")" << "\n";
4975 
4976  for(std::size_t i=0; i<expr.operands().size(); i++)
4977  {
4978  out << "(assert (= (select " << id << " ";
4979  convert_expr(from_integer(i, array_type.index_type()));
4980  out << ") "; // select
4981  if(array_type.element_type().id() == ID_bool && !use_array_of_bool)
4982  {
4983  out << "(ite ";
4984  convert_expr(expr.operands()[i]);
4985  out << " #b1 #b0)";
4986  }
4987  else
4988  {
4989  convert_expr(expr.operands()[i]);
4990  }
4991  out << "))" << "\n"; // =, assert
4992  }
4993 
4994  defined_expressions[expr]=id;
4995  }
4996  }
4997  else if(expr.id()==ID_string_constant)
4998  {
4999  if(defined_expressions.find(expr)==defined_expressions.end())
5000  {
5001  // introduce a temporary array.
5003  const array_typet &array_type=to_array_type(tmp.type());
5004 
5005  const irep_idt id =
5006  "string." + std::to_string(defined_expressions.size());
5007  out << "; the following is a substitute for a string" << "\n";
5008  out << "(declare-fun " << id << " () ";
5009  convert_type(array_type);
5010  out << ")" << "\n";
5011 
5012  for(std::size_t i=0; i<tmp.operands().size(); i++)
5013  {
5014  out << "(assert (= (select " << id << ' ';
5015  convert_expr(from_integer(i, array_type.index_type()));
5016  out << ") "; // select
5017  convert_expr(tmp.operands()[i]);
5018  out << "))" << "\n";
5019  }
5020 
5021  defined_expressions[expr]=id;
5022  }
5023  }
5024  else if(
5025  const auto object_size = expr_try_dynamic_cast<object_size_exprt>(expr))
5026  {
5027  if(object_sizes.find(*object_size) == object_sizes.end())
5028  {
5029  const irep_idt id = convert_identifier(
5030  "object_size." + std::to_string(object_sizes.size()));
5031  out << "(declare-fun " << id << " () ";
5033  out << ")"
5034  << "\n";
5035 
5036  object_sizes[*object_size] = id;
5037  }
5038  }
5039  // clang-format off
5040  else if(!use_FPA_theory &&
5041  expr.operands().size() >= 1 &&
5042  (expr.id() == ID_floatbv_plus ||
5043  expr.id() == ID_floatbv_minus ||
5044  expr.id() == ID_floatbv_mult ||
5045  expr.id() == ID_floatbv_div ||
5046  expr.id() == ID_floatbv_typecast ||
5047  expr.id() == ID_ieee_float_equal ||
5048  expr.id() == ID_ieee_float_notequal ||
5049  ((expr.id() == ID_lt ||
5050  expr.id() == ID_gt ||
5051  expr.id() == ID_le ||
5052  expr.id() == ID_ge ||
5053  expr.id() == ID_isnan ||
5054  expr.id() == ID_isnormal ||
5055  expr.id() == ID_isfinite ||
5056  expr.id() == ID_isinf ||
5057  expr.id() == ID_sign ||
5058  expr.id() == ID_unary_minus ||
5059  expr.id() == ID_typecast ||
5060  expr.id() == ID_abs) &&
5061  to_multi_ary_expr(expr).op0().type().id() == ID_floatbv)))
5062  // clang-format on
5063  {
5064  irep_idt function =
5065  convert_identifier("float_bv." + expr.id_string() + floatbv_suffix(expr));
5066 
5067  if(bvfp_set.insert(function).second)
5068  {
5069  out << "; this is a model for " << expr.id() << " : "
5070  << type2id(to_multi_ary_expr(expr).op0().type()) << " -> "
5071  << type2id(expr.type()) << "\n"
5072  << "(define-fun " << function << " (";
5073 
5074  for(std::size_t i = 0; i < expr.operands().size(); i++)
5075  {
5076  if(i!=0)
5077  out << " ";
5078  out << "(op" << i << ' ';
5079  convert_type(expr.operands()[i].type());
5080  out << ')';
5081  }
5082 
5083  out << ") ";
5084  convert_type(expr.type()); // return type
5085  out << ' ';
5086 
5087  exprt tmp1=expr;
5088  for(std::size_t i = 0; i < tmp1.operands().size(); i++)
5089  tmp1.operands()[i]=
5090  smt2_symbolt("op"+std::to_string(i), tmp1.operands()[i].type());
5091 
5092  exprt tmp2=float_bv(tmp1);
5093  tmp2=letify(tmp2);
5094  CHECK_RETURN(!tmp2.is_nil());
5095 
5096  convert_expr(tmp2);
5097 
5098  out << ")\n"; // define-fun
5099  }
5100  }
5101 }
5102 
5104 {
5105  const typet &type = expr.type();
5106  PRECONDITION(type.id()==ID_array);
5107 
5108  if(use_datatypes)
5109  {
5110  return true; // always use array theory when we have datatypes
5111  }
5112  else
5113  {
5114  // arrays inside structs get flattened
5115  if(expr.id()==ID_with)
5116  return use_array_theory(to_with_expr(expr).old());
5117  else if(expr.id()==ID_member)
5118  return false;
5119  else
5120  return true;
5121  }
5122 }
5123 
5125 {
5126  if(type.id()==ID_array)
5127  {
5128  const array_typet &array_type=to_array_type(type);
5129  CHECK_RETURN(array_type.size().is_not_nil());
5130 
5131  // we always use array theory for top-level arrays
5132  const typet &subtype = array_type.element_type();
5133 
5134  // Arrays map the index type to the element type.
5135  out << "(Array ";
5136  convert_type(array_type.index_type());
5137  out << " ";
5138 
5139  if(subtype.id()==ID_bool && !use_array_of_bool)
5140  out << "(_ BitVec 1)";
5141  else
5142  convert_type(array_type.element_type());
5143 
5144  out << ")";
5145  }
5146  else if(type.id()==ID_bool)
5147  {
5148  out << "Bool";
5149  }
5150  else if(type.id() == ID_struct || type.id() == ID_struct_tag)
5151  {
5152  if(use_datatypes)
5153  {
5154  out << datatype_map.at(type);
5155  }
5156  else
5157  {
5158  std::size_t width=boolbv_width(type);
5159 
5160  out << "(_ BitVec " << width << ")";
5161  }
5162  }
5163  else if(type.id()==ID_vector)
5164  {
5165  if(use_datatypes)
5166  {
5167  out << datatype_map.at(type);
5168  }
5169  else
5170  {
5171  std::size_t width=boolbv_width(type);
5172 
5173  out << "(_ BitVec " << width << ")";
5174  }
5175  }
5176  else if(type.id()==ID_code)
5177  {
5178  // These may appear in structs.
5179  // We replace this by "Bool" in order to keep the same
5180  // member count.
5181  out << "Bool";
5182  }
5183  else if(type.id() == ID_union || type.id() == ID_union_tag)
5184  {
5185  std::size_t width=boolbv_width(type);
5187  to_union_type(ns.follow(type)).components().empty() || width != 0,
5188  "failed to get width of union");
5189 
5190  out << "(_ BitVec " << width << ")";
5191  }
5192  else if(type.id()==ID_pointer)
5193  {
5194  out << "(_ BitVec "
5195  << boolbv_width(type) << ")";
5196  }
5197  else if(type.id()==ID_bv ||
5198  type.id()==ID_fixedbv ||
5199  type.id()==ID_unsignedbv ||
5200  type.id()==ID_signedbv ||
5201  type.id()==ID_c_bool)
5202  {
5203  out << "(_ BitVec "
5204  << to_bitvector_type(type).get_width() << ")";
5205  }
5206  else if(type.id()==ID_c_enum)
5207  {
5208  // these have an underlying type
5209  out << "(_ BitVec "
5210  << to_bitvector_type(to_c_enum_type(type).underlying_type()).get_width()
5211  << ")";
5212  }
5213  else if(type.id()==ID_c_enum_tag)
5214  {
5216  }
5217  else if(type.id()==ID_floatbv)
5218  {
5219  const floatbv_typet &floatbv_type=to_floatbv_type(type);
5220 
5221  if(use_FPA_theory)
5222  out << "(_ FloatingPoint "
5223  << floatbv_type.get_e() << " "
5224  << floatbv_type.get_f() + 1 << ")";
5225  else
5226  out << "(_ BitVec "
5227  << floatbv_type.get_width() << ")";
5228  }
5229  else if(type.id()==ID_rational ||
5230  type.id()==ID_real)
5231  out << "Real";
5232  else if(type.id()==ID_integer)
5233  out << "Int";
5234  else if(type.id()==ID_complex)
5235  {
5236  if(use_datatypes)
5237  {
5238  out << datatype_map.at(type);
5239  }
5240  else
5241  {
5242  std::size_t width=boolbv_width(type);
5243 
5244  out << "(_ BitVec " << width << ")";
5245  }
5246  }
5247  else if(type.id()==ID_c_bit_field)
5248  {
5250  }
5251  else
5252  {
5253  UNEXPECTEDCASE("unsupported type: "+type.id_string());
5254  }
5255 }
5256 
5258 {
5259  std::set<irep_idt> recstack;
5260  find_symbols_rec(type, recstack);
5261 }
5262 
5264  const typet &type,
5265  std::set<irep_idt> &recstack)
5266 {
5267  if(type.id()==ID_array)
5268  {
5269  const array_typet &array_type=to_array_type(type);
5270  find_symbols(array_type.size());
5271  find_symbols_rec(array_type.element_type(), recstack);
5272  }
5273  else if(type.id()==ID_complex)
5274  {
5275  find_symbols_rec(to_complex_type(type).subtype(), recstack);
5276 
5277  if(use_datatypes &&
5278  datatype_map.find(type)==datatype_map.end())
5279  {
5280  const std::string smt_typename =
5281  "complex." + std::to_string(datatype_map.size());
5282  datatype_map[type] = smt_typename;
5283 
5284  out << "(declare-datatypes ((" << smt_typename << " 0)) "
5285  << "(((mk-" << smt_typename;
5286 
5287  out << " (" << smt_typename << ".imag ";
5288  convert_type(to_complex_type(type).subtype());
5289  out << ")";
5290 
5291  out << " (" << smt_typename << ".real ";
5292  convert_type(to_complex_type(type).subtype());
5293  out << ")";
5294 
5295  out << "))))\n";
5296  }
5297  }
5298  else if(type.id()==ID_vector)
5299  {
5300  find_symbols_rec(to_vector_type(type).element_type(), recstack);
5301 
5302  if(use_datatypes &&
5303  datatype_map.find(type)==datatype_map.end())
5304  {
5305  const vector_typet &vector_type=to_vector_type(type);
5306 
5307  mp_integer size = numeric_cast_v<mp_integer>(vector_type.size());
5308 
5309  const std::string smt_typename =
5310  "vector." + std::to_string(datatype_map.size());
5311  datatype_map[type] = smt_typename;
5312 
5313  out << "(declare-datatypes ((" << smt_typename << " 0)) "
5314  << "(((mk-" << smt_typename;
5315 
5316  for(mp_integer i=0; i!=size; ++i)
5317  {
5318  out << " (" << smt_typename << "." << i << " ";
5319  convert_type(to_vector_type(type).element_type());
5320  out << ")";
5321  }
5322 
5323  out << "))))\n";
5324  }
5325  }
5326  else if(type.id() == ID_struct)
5327  {
5328  // Cater for mutually recursive struct types
5329  bool need_decl=false;
5330  if(use_datatypes &&
5331  datatype_map.find(type)==datatype_map.end())
5332  {
5333  const std::string smt_typename =
5334  "struct." + std::to_string(datatype_map.size());
5335  datatype_map[type] = smt_typename;
5336  need_decl=true;
5337  }
5338 
5339  const struct_typet::componentst &components =
5340  to_struct_type(type).components();
5341 
5342  for(const auto &component : components)
5343  find_symbols_rec(component.type(), recstack);
5344 
5345  // Declare the corresponding SMT type if we haven't already.
5346  if(need_decl)
5347  {
5348  const std::string &smt_typename = datatype_map.at(type);
5349 
5350  // We're going to create a datatype named something like `struct.0'.
5351  // It's going to have a single constructor named `mk-struct.0' with an
5352  // argument for each member of the struct. The declaration that
5353  // creates this type looks like:
5354  //
5355  // (declare-datatypes ((struct.0 0)) (((mk-struct.0
5356  // (struct.0.component1 type1)
5357  // ...
5358  // (struct.0.componentN typeN)))))
5359  out << "(declare-datatypes ((" << smt_typename << " 0)) "
5360  << "(((mk-" << smt_typename << " ";
5361 
5362  for(const auto &component : components)
5363  {
5364  out << "(" << smt_typename << "." << component.get_name()
5365  << " ";
5366  convert_type(component.type());
5367  out << ") ";
5368  }
5369 
5370  out << "))))" << "\n";
5371 
5372  // Let's also declare convenience functions to update individual
5373  // members of the struct whil we're at it. The functions are
5374  // named like `update-struct.0.component1'. Their declarations
5375  // look like:
5376  //
5377  // (declare-fun update-struct.0.component1
5378  // ((s struct.0) ; first arg -- the struct to update
5379  // (v type1)) ; second arg -- the value to update
5380  // struct.0 ; the output type
5381  // (mk-struct.0 ; build the new struct...
5382  // v ; the updated value
5383  // (struct.0.component2 s) ; retain the other members
5384  // ...
5385  // (struct.0.componentN s)))
5386 
5387  for(struct_union_typet::componentst::const_iterator
5388  it=components.begin();
5389  it!=components.end();
5390  ++it)
5391  {
5393  out << "(define-fun update-" << smt_typename << "."
5394  << component.get_name() << " "
5395  << "((s " << smt_typename << ") "
5396  << "(v ";
5397  convert_type(component.type());
5398  out << ")) " << smt_typename << " "
5399  << "(mk-" << smt_typename
5400  << " ";
5401 
5402  for(struct_union_typet::componentst::const_iterator
5403  it2=components.begin();
5404  it2!=components.end();
5405  ++it2)
5406  {
5407  if(it==it2)
5408  out << "v ";
5409  else
5410  {
5411  out << "(" << smt_typename << "."
5412  << it2->get_name() << " s) ";
5413  }
5414  }
5415 
5416  out << "))" << "\n";
5417  }
5418 
5419  out << "\n";
5420  }
5421  }
5422  else if(type.id() == ID_union)
5423  {
5424  const union_typet::componentst &components =
5425  to_union_type(type).components();
5426 
5427  for(const auto &component : components)
5428  find_symbols_rec(component.type(), recstack);
5429  }
5430  else if(type.id()==ID_code)
5431  {
5432  const code_typet::parameterst &parameters=
5433  to_code_type(type).parameters();
5434  for(const auto &param : parameters)
5435  find_symbols_rec(param.type(), recstack);
5436 
5437  find_symbols_rec(to_code_type(type).return_type(), recstack);
5438  }
5439  else if(type.id()==ID_pointer)
5440  {
5441  find_symbols_rec(to_pointer_type(type).base_type(), recstack);
5442  }
5443  else if(type.id() == ID_struct_tag)
5444  {
5445  const auto &struct_tag = to_struct_tag_type(type);
5446  const irep_idt &id = struct_tag.get_identifier();
5447 
5448  if(recstack.find(id) == recstack.end())
5449  {
5450  const auto &base_struct = ns.follow_tag(struct_tag);
5451  recstack.insert(id);
5452  find_symbols_rec(base_struct, recstack);
5453  datatype_map[type] = datatype_map[base_struct];
5454  }
5455  }
5456  else if(type.id() == ID_union_tag)
5457  {
5458  const auto &union_tag = to_union_tag_type(type);
5459  const irep_idt &id = union_tag.get_identifier();
5460 
5461  if(recstack.find(id) == recstack.end())
5462  {
5463  recstack.insert(id);
5464  find_symbols_rec(ns.follow_tag(union_tag), recstack);
5465  }
5466  }
5467  else if(type.id() == ID_mathematical_function)
5468  {
5469  const auto &mathematical_function_type =
5471  for(auto &d_type : mathematical_function_type.domain())
5472  find_symbols_rec(d_type, recstack);
5473 
5474  find_symbols_rec(mathematical_function_type.codomain(), recstack);
5475  }
5476 }
5477 
5479 {
5480  return number_of_solver_calls;
5481 }
with_exprt
Operator to update elements in structs and arrays.
Definition: std_expr.h:2423
smt2_convt::set_to
void set_to(const exprt &expr, bool value) override
For a Boolean expression expr, add the constraint 'expr' if value is true, otherwise add 'not expr'.
Definition: smt2_conv.cpp:4609
to_union_tag_type
const union_tag_typet & to_union_tag_type(const typet &type)
Cast a typet to a union_tag_typet.
Definition: c_types.h:202
smt2_convt::convert_typecast
void convert_typecast(const typecast_exprt &expr)
Definition: smt2_conv.cpp:2343
exprt::copy_to_operands
void copy_to_operands(const exprt &expr)
Copy the given argument to the end of exprt's operands.
Definition: expr.h:155
smt2_convt::solvert::CVC4
@ CVC4
smt2_convt::convert_euclidean_mod
void convert_euclidean_mod(const euclidean_mod_exprt &expr)
Definition: smt2_conv.cpp:3290
UNREACHABLE
#define UNREACHABLE
This should be used to mark dead code.
Definition: invariant.h:503
struct_union_typet::components
const componentst & components() const
Definition: std_types.h:147
dstringt
dstringt has one field, an unsigned integer no which is an index into a static table of strings.
Definition: dstring.h:36
pointer_offset_size.h
UNEXPECTEDCASE
#define UNEXPECTEDCASE(S)
Definition: smt2_conv.cpp:50
to_union_type
const union_typet & to_union_type(const typet &type)
Cast a typet to a union_typet.
Definition: c_types.h:162
smt2_convt::solver
solvert solver
Definition: smt2_conv.h:92
to_unsignedbv_type
const unsignedbv_typet & to_unsignedbv_type(const typet &type)
Cast a typet to an unsignedbv_typet.
Definition: bitvector_types.h:191
to_array_comprehension_expr
const array_comprehension_exprt & to_array_comprehension_expr(const exprt &expr)
Cast an exprt to a array_comprehension_exprt.
Definition: std_expr.h:3418
configt::bv_encodingt::object_bits
std::size_t object_bits
Definition: config.h:336
to_vector_expr
const vector_exprt & to_vector_expr(const exprt &expr)
Cast an exprt to an vector_exprt.
Definition: std_expr.h:1692
ieee_floatt
Definition: ieee_float.h:116
format
static format_containert< T > format(const T &o)
Definition: format.h:37
to_extractbit_expr
const extractbit_exprt & to_extractbit_expr(const exprt &expr)
Cast an exprt to an extractbit_exprt.
Definition: bitvector_expr.h:429
smt2_convt::walk_array_tree
void walk_array_tree(std::unordered_map< int64_t, exprt > *operands_map, const irept &src, const array_typet &type)
This function walks the SMT output and populates a map with index/value pairs for the array.
Definition: smt2_conv.cpp:521
to_unary_expr
const unary_exprt & to_unary_expr(const exprt &expr)
Cast an exprt to a unary_exprt.
Definition: std_expr.h:361
smt2_conv.h
literal_exprt::get_literal
literalt get_literal() const
Definition: literal_expr.h:26
UNIMPLEMENTED
#define UNIMPLEMENTED
Definition: invariant.h:525
to_c_bool_type
const c_bool_typet & to_c_bool_type(const typet &type)
Cast a typet to a c_bool_typet.
Definition: c_types.h:106
extractbits_exprt::lower
exprt & lower()
Definition: bitvector_expr.h:479
boolbv_widtht::membert
Definition: boolbv_width.h:39
smt2_convt::solvert::CPROVER_SMT2
@ CPROVER_SMT2
mp_integer
BigInt mp_integer
Definition: smt_terms.h:17
smt2_convt::print_assignment
void print_assignment(std::ostream &out) const override
Print satisfying assignment to out.
Definition: smt2_conv.cpp:131
to_div_expr
const div_exprt & to_div_expr(const exprt &expr)
Cast an exprt to a div_exprt.
Definition: std_expr.h:1146
binary_exprt::rhs
exprt & rhs()
Definition: std_expr.h:623
exprt::depth_end
depth_iteratort depth_end()
Definition: expr.cpp:267
configt::bv_encoding
struct configt::bv_encodingt bv_encoding
arith_tools.h
struct_union_typet::get_component
const componentt & get_component(const irep_idt &component_name) const
Get the reference to a component with given name.
Definition: std_types.cpp:63
smt2_convt::wheret
wheret
Definition: smt2_conv.h:211
smt2_convt::get_number_of_solver_calls
std::size_t get_number_of_solver_calls() const override
Return the number of incremental solver calls.
Definition: smt2_conv.cpp:5478
to_struct_type
const struct_typet & to_struct_type(const typet &type)
Cast a typet to a struct_typet.
Definition: std_types.h:308
address_of_exprt::object
exprt & object()
Definition: pointer_expr.h:380
smt2_convt::wheret::END
@ END
smt2_convt::letify
letifyt letify
Definition: smt2_conv.h:158
float_bv
exprt float_bv(const exprt &src)
Definition: float_bv.h:187
smt2_convt::convert_type
void convert_type(const typet &)
Definition: smt2_conv.cpp:5124
nondet_symbol_exprt::get_identifier
const irep_idt & get_identifier() const
Definition: std_expr.h:270
pos
literalt pos(literalt a)
Definition: literal.h:194
to_euclidean_mod_expr
const euclidean_mod_exprt & to_euclidean_mod_expr(const exprt &expr)
Cast an exprt to a euclidean_mod_exprt.
Definition: std_expr.h:1285
c_bit_field_replacement_type.h
CHECK_RETURN
#define CHECK_RETURN(CONDITION)
Definition: invariant.h:495
floatbv_typecast_exprt::op
exprt & op()
Definition: floatbv_expr.h:30
typet
The type of an expression, extends irept.
Definition: type.h:28
code_typet::parameterst
std::vector< parametert > parameterst
Definition: std_types.h:541
irept::pretty
std::string pretty(unsigned indent=0, unsigned max_indent=0) const
Definition: irep.cpp:495
smt2_convt::pop
void pop() override
Currently, only implements a single stack element (no nested contexts)
Definition: smt2_conv.cpp:904
to_byte_extract_expr
const byte_extract_exprt & to_byte_extract_expr(const exprt &expr)
Definition: byte_operators.h:57
smt2_convt::datatype_map
datatype_mapt datatype_map
Definition: smt2_conv.h:245
pointer_logict::get_dynamic_objects
void get_dynamic_objects(std::vector< mp_integer > &objects) const
Definition: pointer_logic.cpp:34
threeval.h
to_floatbv_type
const floatbv_typet & to_floatbv_type(const typet &type)
Cast a typet to a floatbv_typet.
Definition: bitvector_types.h:367
to_index_expr
const index_exprt & to_index_expr(const exprt &expr)
Cast an exprt to an index_exprt.
Definition: std_expr.h:1478
to_if_expr
const if_exprt & to_if_expr(const exprt &expr)
Cast an exprt to an if_exprt.
Definition: std_expr.h:2403
smt2_convt::parse_array
exprt parse_array(const irept &s, const array_typet &type)
This function is for parsing array output from SMT solvers when "(get-value |???|)" returns an array ...
Definition: smt2_conv.cpp:477
smt2_convt::use_array_theory
bool use_array_theory(const exprt &)
Definition: smt2_conv.cpp:5103
smt2_convt::convert_floatbv
void convert_floatbv(const exprt &expr)
Definition: smt2_conv.cpp:1012
floatbv_typet::get_f
std::size_t get_f() const
Definition: bitvector_types.cpp:26
vector_typet::size
const constant_exprt & size() const
Definition: std_types.cpp:269
smt2_convt::use_datatypes
bool use_datatypes
Definition: smt2_conv.h:65
smt2_convt::convert
literalt convert(const exprt &expr)
Definition: smt2_conv.cpp:804
if_exprt
The trinary if-then-else operator.
Definition: std_expr.h:2322
to_isinf_expr
const isinf_exprt & to_isinf_expr(const exprt &expr)
Cast an exprt to a isinf_exprt.
Definition: floatbv_expr.h:157
has_byte_operator
bool has_byte_operator(const exprt &src)
Definition: byte_operators.cpp:2377
floatbv_typet
Fixed-width bit-vector with IEEE floating-point interpretation.
Definition: bitvector_types.h:321
boolbv_widtht::membert::offset
std::size_t offset
Definition: boolbv_width.h:41
pointer_predicates.h
string2integer
const mp_integer string2integer(const std::string &n, unsigned base)
Definition: mp_arith.cpp:54
replication_exprt::times
constant_exprt & times()
Definition: bitvector_expr.h:547
to_c_enum_type
const c_enum_typet & to_c_enum_type(const typet &type)
Cast a typet to a c_enum_typet.
Definition: c_types.h:302
lower_byte_update
static exprt lower_byte_update(const byte_update_exprt &src, const exprt &value_as_byte_array, const optionalt< exprt > &non_const_update_bound, const namespacet &ns)
Apply a byte update src using the byte array value_as_byte_array as update value.
Definition: byte_operators.cpp:2080
fixedbv_spect::get_fraction_bits
std::size_t get_fraction_bits() const
Definition: fixedbv.h:35
prefix.h
with_exprt::new_value
exprt & new_value()
Definition: std_expr.h:2454
has_quantifier
static bool has_quantifier(const exprt &expr)
Definition: smt2_conv.cpp:794
fixedbv.h
invariant.h
smt2_convt::solvert::MATHSAT
@ MATHSAT
pointer_offset_exprt::pointer
exprt & pointer()
Definition: pointer_expr.h:897
smt2_convt::convert_update
void convert_update(const exprt &expr)
Definition: smt2_conv.cpp:4251
s1
int8_t s1
Definition: bytecode_info.h:59
union_exprt
Union constructor from single element.
Definition: std_expr.h:1707
to_string_constant
const string_constantt & to_string_constant(const exprt &expr)
Definition: string_constant.h:40
to_array_of_expr
const array_of_exprt & to_array_of_expr(const exprt &expr)
Cast an exprt to an array_of_exprt.
Definition: std_expr.h:1543
smt2_convt::convert_plus
void convert_plus(const plus_exprt &expr)
Definition: smt2_conv.cpp:3468
plus_exprt
The plus expression Associativity is not specified.
Definition: std_expr.h:946
smt2_convt::convert_floatbv_plus
void convert_floatbv_plus(const ieee_float_op_exprt &expr)
Definition: smt2_conv.cpp:3661
lower_byte_extract
exprt lower_byte_extract(const byte_extract_exprt &src, const namespacet &ns)
rewrite byte extraction from an array to byte extraction from a concatenation of array index expressi...
Definition: byte_operators.cpp:1099
fixedbv_spect::integer_bits
std::size_t integer_bits
Definition: fixedbv.h:22
to_mathematical_function_type
const mathematical_function_typet & to_mathematical_function_type(const typet &type)
Cast a typet to a mathematical_function_typet.
Definition: mathematical_types.h:119
string_constant.h
exprt
Base class for all expressions.
Definition: expr.h:55
binary_exprt::lhs
exprt & lhs()
Definition: std_expr.h:613
struct_union_typet::componentst
std::vector< componentt > componentst
Definition: std_types.h:140
unary_exprt
Generic base class for unary expressions.
Definition: std_expr.h:313
smt2_convt::convert_expr
void convert_expr(const exprt &)
Definition: smt2_conv.cpp:1046
binary_exprt
A base class for binary expressions.
Definition: std_expr.h:582
to_bv_type
const bv_typet & to_bv_type(const typet &type)
Cast a typet to a bv_typet.
Definition: bitvector_types.h:82
to_union_expr
const union_exprt & to_union_expr(const exprt &expr)
Cast an exprt to a union_exprt.
Definition: std_expr.h:1754
namespace_baset::follow_tag
const union_typet & follow_tag(const union_tag_typet &) const
Follow type tag of union type.
Definition: namespace.cpp:63
bswap_exprt::get_bits_per_byte
std::size_t get_bits_per_byte() const
Definition: bitvector_expr.h:33
smt2_convt::convert_constant
void convert_constant(const constant_exprt &expr)
Definition: smt2_conv.cpp:3146
sign_exprt
Sign of an expression Predicate is true if _op is negative, false otherwise.
Definition: std_expr.h:538
exprt::op0
exprt & op0()
Definition: expr.h:125
smt2_convt::convert_literal
void convert_literal(const literalt)
Definition: smt2_conv.cpp:869
component
auto component(T &struct_expr, const irep_idt &name, const namespacet &ns) -> decltype(struct_expr.op0())
Definition: std_expr.cpp:76
smt2_convt::handle
exprt handle(const exprt &expr) override
Generate a handle, which is an expression that has the same value as the argument in any model that i...
Definition: smt2_conv.cpp:860
irep_idt
dstringt irep_idt
Definition: irep.h:37
vector_typet
The vector type.
Definition: std_types.h:1007
object_size_exprt::pointer
exprt & pointer()
Definition: pointer_expr.h:1014
to_integer
bool to_integer(const constant_exprt &expr, mp_integer &int_value)
Convert a constant expression expr to an arbitrary-precision integer.
Definition: arith_tools.cpp:20
bool_typet
The Boolean type.
Definition: std_types.h:35
quantifier_exprt
A base class for quantifier expressions.
Definition: mathematical_expr.h:267
to_string
std::string to_string(const string_not_contains_constraintt &expr)
Used for debug printing.
Definition: string_constraint.cpp:58
to_bitnot_expr
const bitnot_exprt & to_bitnot_expr(const exprt &expr)
Cast an exprt to a bitnot_exprt.
Definition: bitvector_expr.h:106
vector_exprt
Vector constructor from list of elements.
Definition: std_expr.h:1671
exprt::is_true
bool is_true() const
Return whether the expression is a constant representing true.
Definition: expr.cpp:34
literalt::var_no
var_not var_no() const
Definition: literal.h:83
to_bitvector_type
const bitvector_typet & to_bitvector_type(const typet &type)
Cast a typet to a bitvector_typet.
Definition: bitvector_types.h:32
make_binary
exprt make_binary(const exprt &expr)
splits an expression with >=3 operands into nested binary expressions
Definition: expr_util.cpp:36
to_complex_type
const complex_typet & to_complex_type(const typet &type)
Cast a typet to a complex_typet.
Definition: std_types.h:1102
div_exprt
Division.
Definition: std_expr.h:1096
shift_exprt::op
exprt & op()
Definition: bitvector_expr.h:239
pointer_logict::get_invalid_object
const mp_integer & get_invalid_object() const
Definition: pointer_logic.h:62
namespace.h
smt2_convt::solvert::CVC3
@ CVC3
object_size_exprt
Expression for finding the size (in bytes) of the object a pointer points to.
Definition: pointer_expr.h:1006
index_type
bitvector_typet index_type()
Definition: c_types.cpp:22
equal_exprt
Equality.
Definition: std_expr.h:1305
CHECK_RETURN_WITH_DIAGNOSTICS
#define CHECK_RETURN_WITH_DIAGNOSTICS(CONDITION,...)
Definition: invariant.h:496
smt2_convt::type2id
std::string type2id(const typet &) const
Definition: smt2_conv.cpp:961
to_popcount_expr
const popcount_exprt & to_popcount_expr(const exprt &expr)
Cast an exprt to a popcount_exprt.
Definition: bitvector_expr.h:685
smt2_convt::flatten_array
void flatten_array(const exprt &)
produce a flat bit-vector for a given array of fixed size
Definition: smt2_conv.cpp:3086
to_minus_expr
const minus_exprt & to_minus_expr(const exprt &expr)
Cast an exprt to a minus_exprt.
Definition: std_expr.h:1031
to_find_first_set_expr
const find_first_set_exprt & to_find_first_set_expr(const exprt &expr)
Cast an exprt to a find_first_set_exprt.
Definition: bitvector_expr.h:1489
replication_exprt
Bit-vector replication.
Definition: bitvector_expr.h:535
if_exprt::false_case
exprt & false_case()
Definition: std_expr.h:2360
smt2_convt::solvert::BOOLECTOR
@ BOOLECTOR
isfinite_exprt
Evaluates to true if the operand is finite.
Definition: floatbv_expr.h:175
notequal_exprt
Disequality.
Definition: std_expr.h:1364
smt2_convt::convert_is_dynamic_object
void convert_is_dynamic_object(const unary_exprt &)
Definition: smt2_conv.cpp:3324
exprt::is_false
bool is_false() const
Return whether the expression is a constant representing false.
Definition: expr.cpp:43
unsignedbv_typet
Fixed-width bit-vector with unsigned binary interpretation.
Definition: bitvector_types.h:158
smt2_convt::convert_index
void convert_index(const index_exprt &expr)
Definition: smt2_conv.cpp:4258
replication_exprt::op
exprt & op()
Definition: bitvector_expr.h:557
to_floatbv_typecast_expr
const floatbv_typecast_exprt & to_floatbv_typecast_expr(const exprt &expr)
Cast an exprt to a floatbv_typecast_exprt.
Definition: floatbv_expr.h:68
extractbits_exprt::upper
exprt & upper()
Definition: bitvector_expr.h:474
to_literal_expr
const literal_exprt & to_literal_expr(const exprt &expr)
Cast a generic exprt to a literal_exprt.
Definition: literal_expr.h:56
to_fixedbv_type
const fixedbv_typet & to_fixedbv_type(const typet &type)
Cast a typet to a fixedbv_typet.
Definition: bitvector_types.h:305
irept::get
const irep_idt & get(const irep_idt &name) const
Definition: irep.cpp:45
ieee_float_spect
Definition: ieee_float.h:22
expr_lowering.h
to_unary_plus_expr
const unary_plus_exprt & to_unary_plus_expr(const exprt &expr)
Cast an exprt to a unary_plus_exprt.
Definition: std_expr.h:497
to_binary_expr
const binary_exprt & to_binary_expr(const exprt &expr)
Cast an exprt to a binary_exprt.
Definition: std_expr.h:660
binding_exprt::variables
variablest & variables()
Definition: std_expr.h:3073
INVARIANT_WITH_DIAGNOSTICS
#define INVARIANT_WITH_DIAGNOSTICS(CONDITION, REASON,...)
Same as invariant, with one or more diagnostics attached Diagnostics can be of any type that has a sp...
Definition: invariant.h:437
to_bswap_expr
const bswap_exprt & to_bswap_expr(const exprt &expr)
Cast an exprt to a bswap_exprt.
Definition: bitvector_expr.h:63
euclidean_mod_exprt
Boute's Euclidean definition of Modulo – to match SMT-LIB2.
Definition: std_expr.h:1235
smt2_convt::convert_rounding_mode_FPA
void convert_rounding_mode_FPA(const exprt &expr)
Converting a constant or symbolic rounding mode to SMT-LIB.
Definition: smt2_conv.cpp:3604
struct_exprt
Struct constructor from list of elements.
Definition: std_expr.h:1818
array_typet::size
const exprt & size() const
Definition: std_types.h:800
namespacet
A namespacet is essentially one or two symbol tables bound together, to allow for symbol lookups in t...
Definition: namespace.h:90
can_cast_expr< overflow_result_exprt >
bool can_cast_expr< overflow_result_exprt >(const exprt &base)
Definition: bitvector_expr.h:1387
to_nondet_symbol_expr
const nondet_symbol_exprt & to_nondet_symbol_expr(const exprt &expr)
Cast an exprt to a nondet_symbol_exprt.
Definition: std_expr.h:293
let_exprt::variables
binding_exprt::variablest & variables()
convenience accessor for binding().variables()
Definition: std_expr.h:3223
string2int.h
smt2_convt::notes
std::string notes
Definition: smt2_conv.h:91
exprt::type
typet & type()
Return the type of the expression.
Definition: expr.h:84
literalt::is_true
bool is_true() const
Definition: literal.h:156
to_pointer_object_expr
const pointer_object_exprt & to_pointer_object_expr(const exprt &expr)
Cast an exprt to a pointer_object_exprt.
Definition: pointer_expr.h:986
smt2_convt::to_smt2_symbol
const smt2_symbolt & to_smt2_symbol(const exprt &expr)
Definition: smt2_conv.h:201
irept::is_not_nil
bool is_not_nil() const
Definition: irep.h:380
smt2_convt::number_of_solver_calls
std::size_t number_of_solver_calls
Definition: smt2_conv.h:97
with_exprt::old
exprt & old()
Definition: std_expr.h:2434
smt2_convt::convert_mult
void convert_mult(const mult_exprt &expr)
Definition: smt2_conv.cpp:3895
expr_protectedt::op0
exprt & op0()
Definition: expr.h:125
byte_operators.h
Expression classes for byte-level operators.
to_code_type
const code_typet & to_code_type(const typet &type)
Cast a typet to a code_typet.
Definition: std_types.h:744
smt2_convt::convert_identifier
static std::string convert_identifier(const irep_idt &identifier)
Definition: smt2_conv.cpp:926
smt2_convt::convert_floatbv_div
void convert_floatbv_div(const ieee_float_op_exprt &expr)
Definition: smt2_conv.cpp:3875
vector_typet::element_type
const typet & element_type() const
The type of the elements of the vector.
Definition: std_types.h:1026
smt2_convt::prepare_for_convert_expr
exprt prepare_for_convert_expr(const exprt &expr)
Perform steps necessary before an expression is passed to convert_expr.
Definition: smt2_conv.cpp:4785
smt2_convt::push
void push() override
Unimplemented.
Definition: smt2_conv.cpp:892
smt2_convt::set_values
std::unordered_map< irep_idt, bool > set_values
The values which boolean identifiers have been smt2_convt::set_to or in other words those which are a...
Definition: smt2_conv.h:258
smt2_convt::parse_literal
constant_exprt parse_literal(const irept &, const typet &type)
Definition: smt2_conv.cpp:339
DATA_INVARIANT_WITH_DIAGNOSTICS
#define DATA_INVARIANT_WITH_DIAGNOSTICS(CONDITION, REASON,...)
Definition: invariant.h:511
has_prefix
bool has_prefix(const std::string &s, const std::string &prefix)
Definition: converter.cpp:13
integer2bvrep
irep_idt integer2bvrep(const mp_integer &src, std::size_t width)
convert an integer to bit-vector representation with given width This uses two's complement for negat...
Definition: arith_tools.cpp:380
to_mod_expr
const mod_exprt & to_mod_expr(const exprt &expr)
Cast an exprt to a mod_exprt.
Definition: std_expr.h:1217
ieee_float_spect::width
std::size_t width() const
Definition: ieee_float.h:51
smt2_convt::solvert::YICES
@ YICES
DATA_INVARIANT
#define DATA_INVARIANT(CONDITION, REASON)
This condition should be used to document that assumptions that are made on goto_functions,...
Definition: invariant.h:510
to_mult_expr
const mult_exprt & to_mult_expr(const exprt &expr)
Cast an exprt to a mult_exprt.
Definition: std_expr.h:1077
smt2_convt::write_footer
void write_footer()
Writes the end of the SMT file to the smt_convt::out stream.
Definition: smt2_conv.cpp:184
id2string
const std::string & id2string(const irep_idt &d)
Definition: irep.h:47
struct_union_typet::componentt::get_name
const irep_idt & get_name() const
Definition: std_types.h:79
smt2_convt::convert_relation
void convert_relation(const binary_relation_exprt &)
Definition: smt2_conv.cpp:3361
floatbv_typecast_exprt
Semantic type conversion from/to floating-point formats.
Definition: floatbv_expr.h:18
smt2_convt::identifier_map
identifier_mapt identifier_map
Definition: smt2_conv.h:238
get_identifier
static optionalt< smt_termt > get_identifier(const exprt &expr, const std::unordered_map< exprt, smt_identifier_termt, irep_hash > &expression_handle_identifiers, const std::unordered_map< exprt, smt_identifier_termt, irep_hash > &expression_identifiers)
Definition: smt2_incremental_decision_procedure.cpp:328
to_c_enum_tag_type
const c_enum_tag_typet & to_c_enum_tag_type(const typet &type)
Cast a typet to a c_enum_tag_typet.
Definition: c_types.h:344
forall_operands
#define forall_operands(it, expr)
Definition: expr.h:20
to_byte_update_expr
const byte_update_exprt & to_byte_update_expr(const exprt &expr)
Definition: byte_operators.h:117
pointer_logict::pointert::offset
mp_integer offset
Definition: pointer_logic.h:30
smt2_convt::convert_mod
void convert_mod(const mod_exprt &expr)
Definition: smt2_conv.cpp:3305
member_exprt::struct_op
const exprt & struct_op() const
Definition: std_expr.h:2832
exprt::visit_post
void visit_post(std::function< void(exprt &)>)
These are post-order traversal visitors, i.e., the visitor is executed on a node after its children h...
Definition: expr.cpp:216
smt2_tokenizer.h
multi_ary_exprt::op1
exprt & op1()
Definition: std_expr.h:883
PRECONDITION
#define PRECONDITION(CONDITION)
Definition: invariant.h:463
symbol_exprt::get_identifier
const irep_idt & get_identifier() const
Definition: std_expr.h:142
literal_exprt
Definition: literal_expr.h:17
nil_exprt
The NIL expression.
Definition: std_expr.h:3025
can_cast_expr< plus_overflow_exprt >
bool can_cast_expr< plus_overflow_exprt >(const exprt &base)
Definition: bitvector_expr.h:811
to_isfinite_expr
const isfinite_exprt & to_isfinite_expr(const exprt &expr)
Cast an exprt to a isfinite_exprt.
Definition: floatbv_expr.h:201
smt2_convt::smt2_symbolt::get_identifier
const irep_idt & get_identifier() const
Definition: smt2_conv.h:195
to_let_expr
const let_exprt & to_let_expr(const exprt &expr)
Cast an exprt to a let_exprt.
Definition: std_expr.h:3266
to_plus_expr
const plus_exprt & to_plus_expr(const exprt &expr)
Cast an exprt to a plus_exprt.
Definition: std_expr.h:986
smt2_convt::convert_floatbv_mult
void convert_floatbv_mult(const ieee_float_op_exprt &expr)
Definition: smt2_conv.cpp:3970
fixedbv_spect::width
std::size_t width
Definition: fixedbv.h:22
extractbits_exprt::src
exprt & src()
Definition: bitvector_expr.h:469
smt2_convt::use_check_sat_assuming
bool use_check_sat_assuming
Definition: smt2_conv.h:64
literalt::is_false
bool is_false() const
Definition: literal.h:161
irept::id_string
const std::string & id_string() const
Definition: irep.h:399
exprt::op1
exprt & op1()
Definition: expr.h:128
const_literal
literalt const_literal(bool value)
Definition: literal.h:188
to_notequal_expr
const notequal_exprt & to_notequal_expr(const exprt &expr)
Cast an exprt to an notequal_exprt.
Definition: std_expr.h:1390
mult_exprt
Binary multiplication Associativity is not specified.
Definition: std_expr.h:1051
floatbv_typet::get_e
std::size_t get_e() const
Definition: bitvector_types.h:328
to_pointer_type
const pointer_typet & to_pointer_type(const typet &type)
Cast a typet to a pointer_typet.
Definition: pointer_expr.h:79
simplify_expr
exprt simplify_expr(exprt src, const namespacet &ns)
Definition: simplify_expr.cpp:2931
index_exprt::index
exprt & index()
Definition: std_expr.h:1450
to_unary_minus_expr
const unary_minus_exprt & to_unary_minus_expr(const exprt &expr)
Cast an exprt to a unary_minus_exprt.
Definition: std_expr.h:453
ieee_floatt::pack
mp_integer pack() const
Definition: ieee_float.cpp:374
smt2_convt::convert_struct
void convert_struct(const struct_exprt &expr)
Definition: smt2_conv.cpp:3020
pointer_type
pointer_typet pointer_type(const typet &subtype)
Definition: c_types.cpp:253
unary_minus_exprt
The unary minus expression.
Definition: std_expr.h:422
smt2_convt::define_object_size
void define_object_size(const irep_idt &id, const object_size_exprt &expr)
Definition: smt2_conv.cpp:235
index_exprt::array
exprt & array()
Definition: std_expr.h:1440
irept::swap
void swap(irept &irep)
Definition: irep.h:442
extractbit_exprt::index
exprt & index()
Definition: bitvector_expr.h:396
to_symbol_expr
const symbol_exprt & to_symbol_expr(const exprt &expr)
Cast an exprt to a symbol_exprt.
Definition: std_expr.h:222
smt2_convt::use_lambda_for_array
bool use_lambda_for_array
Definition: smt2_conv.h:66
struct_union_typet::has_component
bool has_component(const irep_idt &component_name) const
Definition: std_types.h:157
smt2_convt::out
std::ostream & out
Definition: smt2_conv.h:90
object_size
exprt object_size(const exprt &pointer)
Definition: pointer_predicates.cpp:33
decision_proceduret::resultt::D_ERROR
@ D_ERROR
irept::is_nil
bool is_nil() const
Definition: irep.h:376
irept::id
const irep_idt & id() const
Definition: irep.h:396
to_struct_tag_type
const struct_tag_typet & to_struct_tag_type(const typet &type)
Cast a typet to a struct_tag_typet.
Definition: std_types.h:474
range.h
exprt::operandst
std::vector< exprt > operandst
Definition: expr.h:58
dstringt::empty
bool empty() const
Definition: dstring.h:88
smt2_convt::bvfp_set
std::set< irep_idt > bvfp_set
Definition: smt2_conv.h:186
abs_exprt
Absolute value.
Definition: std_expr.h:378
to_count_trailing_zeros_expr
const count_trailing_zeros_exprt & to_count_trailing_zeros_expr(const exprt &expr)
Cast an exprt to a count_trailing_zeros_exprt.
Definition: bitvector_expr.h:1136
false_exprt
The Boolean constant false.
Definition: std_expr.h:3016
floatbv_expr.h
union_typet
The union type.
Definition: c_types.h:124
smt2_convt::convert_div
void convert_div(const div_exprt &expr)
Definition: smt2_conv.cpp:3831
unary_exprt::op
const exprt & op() const
Definition: std_expr.h:326
annotated_pointer_constant_exprt
Pointer-typed bitvector constant annotated with the pointer expression that the bitvector is the nume...
Definition: pointer_expr.h:817
smt2_convt::use_array_of_bool
bool use_array_of_bool
Definition: smt2_conv.h:62
code_typet::parameters
const parameterst & parameters() const
Definition: std_types.h:655
ieee_floatt::is_infinity
bool is_infinity() const
Definition: ieee_float.h:250
smt2_convt::convert_minus
void convert_minus(const minus_exprt &expr)
Definition: smt2_conv.cpp:3702
pointer_offset
exprt pointer_offset(const exprt &pointer)
Definition: pointer_predicates.cpp:38
to_shift_expr
const shift_exprt & to_shift_expr(const exprt &expr)
Cast an exprt to a shift_exprt.
Definition: bitvector_expr.h:278
fixedbv_spect
Definition: fixedbv.h:19
tvt
Definition: threeval.h:19
literalt::sign
bool sign() const
Definition: literal.h:88
pointer_logict::objects
numberingt< exprt, irep_hash > objects
Definition: pointer_logic.h:26
decision_proceduret::resultt
resultt
Result of running the decision procedure.
Definition: decision_procedure.h:43
minus_exprt
Binary minus.
Definition: std_expr.h:1005
to_with_expr
const with_exprt & to_with_expr(const exprt &expr)
Cast an exprt to a with_exprt.
Definition: std_expr.h:2486
ieee_floatt::plus_infinity
static ieee_floatt plus_infinity(const ieee_float_spect &_spec)
Definition: ieee_float.h:212
pointer_logict::pointert::object
mp_integer object
Definition: pointer_logic.h:30
config
configt config
Definition: config.cpp:25
to_signedbv_type
const signedbv_typet & to_signedbv_type(const typet &type)
Cast a typet to a signedbv_typet.
Definition: bitvector_types.h:241
smt2_convt::boolbv_width
boolbv_widtht boolbv_width
Definition: smt2_conv.h:95
ieee_floatt::spec
ieee_float_spect spec
Definition: ieee_float.h:135
fixedbv_typet
Fixed-width bit-vector with signed fixed-point interpretation.
Definition: bitvector_types.h:260
smt2_convt::find_symbols_rec
void find_symbols_rec(const typet &type, std::set< irep_idt > &recstack)
Definition: smt2_conv.cpp:5263
simplify_expr.h
bitvector_typet::get_width
std::size_t get_width() const
Definition: std_types.h:876
to_ieee_float_op_expr
const ieee_float_op_exprt & to_ieee_float_op_expr(const exprt &expr)
Cast an exprt to an ieee_float_op_exprt.
Definition: floatbv_expr.h:425
bitnot_exprt
Bit-wise negation of bit-vectors.
Definition: bitvector_expr.h:81
extractbit_exprt
Extracts a single bit of a bit-vector operand.
Definition: bitvector_expr.h:380
exprt::is_zero
bool is_zero() const
Return whether the expression is a constant representing 0.
Definition: expr.cpp:65
smt2_convt::convert_with
void convert_with(const with_exprt &expr)
Definition: smt2_conv.cpp:4013
member_exprt
Extract member of struct or union.
Definition: std_expr.h:2793
struct_union_typet::componentt
Definition: std_types.h:68
to_c_bit_field_type
const c_bit_field_typet & to_c_bit_field_type(const typet &type)
Cast a typet to a c_bit_field_typet.
Definition: c_types.h:58
expr_util.h
Deprecated expression utility functions.
bvrep2integer
mp_integer bvrep2integer(const irep_idt &src, std::size_t width, bool is_signed)
convert a bit-vector representation (possibly signed) to integer
Definition: arith_tools.cpp:402
shift_exprt
A base class for shift and rotate operators.
Definition: bitvector_expr.h:229
format_expr.h
string_constantt::to_array_expr
array_exprt to_array_expr() const
convert string into array constant
Definition: string_constant.cpp:29
smt2_convt::use_FPA_theory
bool use_FPA_theory
Definition: smt2_conv.h:61
smt2_convt::write_header
void write_header()
Definition: smt2_conv.cpp:154
ieee_float_op_exprt::rounding_mode
exprt & rounding_mode()
Definition: floatbv_expr.h:395
struct_typet
Structure type, corresponds to C style structs.
Definition: std_types.h:230
smt2_convt::logic
std::string logic
Definition: smt2_conv.h:91
pointer_offset_size
optionalt< mp_integer > pointer_offset_size(const typet &type, const namespacet &ns)
Compute the size of a type in bytes, rounding up to full bytes.
Definition: pointer_offset_size.cpp:90
array_typet
Arrays with given size.
Definition: std_types.h:762
expr_iterator.h
if_exprt::true_case
exprt & true_case()
Definition: std_expr.h:2350
binding_exprt::where
exprt & where()
Definition: std_expr.h:3083
fixedbv_typet::get_fraction_bits
std::size_t get_fraction_bits() const
Definition: bitvector_types.h:267
smt2_convt::smt2_symbolt
Definition: smt2_conv.h:188
to_quantifier_expr
const quantifier_exprt & to_quantifier_expr(const exprt &expr)
Cast an exprt to a quantifier_exprt.
Definition: mathematical_expr.h:319
boolbv_width.h
namespace_baset::follow
const typet & follow(const typet &) const
Resolve type symbol to the type it points to.
Definition: namespace.cpp:49
to_sign_expr
const sign_exprt & to_sign_expr(const exprt &expr)
Cast an exprt to a sign_exprt.
Definition: std_expr.h:564
to_not_expr
const not_exprt & to_not_expr(const exprt &expr)
Cast an exprt to an not_exprt.
Definition: std_expr.h:2303
shift_exprt::distance
exprt & distance()
Definition: bitvector_expr.h:249
smt2_convt::parse_rec
exprt parse_rec(const irept &s, const typet &type)
Definition: smt2_conv.cpp:636
to_bitreverse_expr
const bitreverse_exprt & to_bitreverse_expr(const exprt &expr)
Cast an exprt to a bitreverse_exprt.
Definition: bitvector_expr.h:1186
isinf_exprt
Evaluates to true if the operand is infinite.
Definition: floatbv_expr.h:131
smt2_convt::dec_solve
resultt dec_solve() override
Run the decision procedure to solve the problem.
Definition: smt2_conv.cpp:272
to_implies_expr
const implies_exprt & to_implies_expr(const exprt &expr)
Cast an exprt to a implies_exprt.
Definition: std_expr.h:2159
smt2_convt::floatbv_suffix
std::string floatbv_suffix(const exprt &) const
Definition: smt2_conv.cpp:1005
exprt::depth_begin
depth_iteratort depth_begin()
Definition: expr.cpp:265
smt2_convt::smt2_convt
smt2_convt(const namespacet &_ns, const std::string &_benchmark, const std::string &_notes, const std::string &_logic, solvert _solver, std::ostream &_out)
Definition: smt2_conv.cpp:55
extractbit_exprt::src
exprt & src()
Definition: bitvector_expr.h:391
from_integer
constant_exprt from_integer(const mp_integer &int_value, const typet &type)
Definition: arith_tools.cpp:100
to_typecast_expr
const typecast_exprt & to_typecast_expr(const exprt &expr)
Cast an exprt to a typecast_exprt.
Definition: std_expr.h:2051
solver
int solver(std::istream &in)
Definition: smt2_solver.cpp:407
ieee_float.h
smt2_convt::assumptions
std::vector< exprt > assumptions
Definition: smt2_conv.h:94
fixedbv_typet::get_integer_bits
std::size_t get_integer_bits() const
Definition: bitvector_types.cpp:19
smt2_convt::lower_byte_operators
exprt lower_byte_operators(const exprt &expr)
Lower byte_update and byte_extract operations within expr.
Definition: smt2_conv.cpp:4753
exprt::is_constant
bool is_constant() const
Return whether the expression is a constant.
Definition: expr.cpp:27
if_exprt::cond
exprt & cond()
Definition: std_expr.h:2340
binary_relation_exprt
A base class for relations, i.e., binary predicates whose two operands have the same type.
Definition: std_expr.h:706
smt2_convt::convert_floatbv_rem
void convert_floatbv_rem(const binary_exprt &expr)
Definition: smt2_conv.cpp:3990
smt2_convt::parse_struct
struct_exprt parse_struct(const irept &s, const struct_typet &type)
Definition: smt2_conv.cpp:577
to_isnormal_expr
const isnormal_exprt & to_isnormal_expr(const exprt &expr)
Cast an exprt to a isnormal_exprt.
Definition: floatbv_expr.h:245
pointer_typet::base_type
const typet & base_type() const
The type of the data what we point to.
Definition: pointer_expr.h:35
smt2_convt::no_boolean_variables
std::size_t no_boolean_variables
Definition: smt2_conv.h:266
literalt
Definition: literal.h:25
smt2_convt::convert_union
void convert_union(const union_exprt &expr)
Definition: smt2_conv.cpp:3117
ieee_float_op_exprt::lhs
exprt & lhs()
Definition: floatbv_expr.h:375
irept::get_sub
subt & get_sub()
Definition: irep.h:456
binary
static std::string binary(const constant_exprt &src)
Definition: json_expr.cpp:189
to_array_type
const array_typet & to_array_type(const typet &type)
Cast a typet to an array_typet.
Definition: std_types.h:844
smt2_convt::ns
const namespacet & ns
Definition: smt2_conv.h:89
same_object
exprt same_object(const exprt &p1, const exprt &p2)
Definition: pointer_predicates.cpp:28
power
mp_integer power(const mp_integer &base, const mp_integer &exponent)
A multi-precision implementation of the power operator.
Definition: arith_tools.cpp:195
to_equal_expr
const equal_exprt & to_equal_expr(const exprt &expr)
Cast an exprt to an equal_exprt.
Definition: std_expr.h:1347
c_bit_field_replacement_type
typet c_bit_field_replacement_type(const c_bit_field_typet &src, const namespacet &ns)
Definition: c_bit_field_replacement_type.cpp:15
smt2_convt::solvert::GENERIC
@ GENERIC
smt2_convt::emit_set_logic
bool emit_set_logic
Definition: smt2_conv.h:67
smt2_convt::pointer_logic
pointer_logict pointer_logic
Definition: smt2_conv.h:216
unsafe_string2size_t
std::size_t unsafe_string2size_t(const std::string &str, int base)
Definition: string2int.cpp:40
can_cast_expr< mult_overflow_exprt >
bool can_cast_expr< mult_overflow_exprt >(const exprt &base)
Definition: bitvector_expr.h:849
smt2_convt::convert_address_of_rec
void convert_address_of_rec(const exprt &expr, const pointer_typet &result_type)
Definition: smt2_conv.cpp:702
config.h
smt2_convt::get
exprt get(const exprt &expr) const override
Return expr with variables replaced by values from satisfying assignment if available.
Definition: smt2_conv.cpp:279
to_pointer_offset_expr
const pointer_offset_exprt & to_pointer_offset_expr(const exprt &expr)
Cast an exprt to a pointer_offset_exprt.
Definition: pointer_expr.h:928
to_extractbits_expr
const extractbits_exprt & to_extractbits_expr(const exprt &expr)
Cast an exprt to an extractbits_exprt.
Definition: bitvector_expr.h:517
to_vector_type
const vector_typet & to_vector_type(const typet &type)
Cast a typet to a vector_typet.
Definition: std_types.h:1060
ieee_float_op_exprt
IEEE floating-point operations These have two data operands (op0 and op1) and one rounding mode (op2)...
Definition: floatbv_expr.h:363
size_of_expr
optionalt< exprt > size_of_expr(const typet &type, const namespacet &ns)
Definition: pointer_offset_size.cpp:280
is_smt2_simple_symbol_character
bool is_smt2_simple_symbol_character(char ch)
Definition: smt2_tokenizer.cpp:11
to_member_expr
const member_exprt & to_member_expr(const exprt &expr)
Cast an exprt to a member_exprt.
Definition: std_expr.h:2886
smt2_convt::solvert::Z3
@ Z3
bswap_exprt
The byte swap expression.
Definition: bitvector_expr.h:18
pointer_object_exprt::pointer
exprt & pointer()
Definition: pointer_expr.h:955
member_exprt::get_component_name
irep_idt get_component_name() const
Definition: std_expr.h:2816
irept
There are a large number of kinds of tree structured or tree-like data in CPROVER.
Definition: irep.h:359
smt2_convt::l_get
tvt l_get(literalt l) const
Definition: smt2_conv.cpp:141
implies_exprt
Boolean implication.
Definition: std_expr.h:2133
ieee_floatt::to_expr
constant_exprt to_expr() const
Definition: ieee_float.cpp:702
ieee_floatt::get_sign
bool get_sign() const
Definition: ieee_float.h:248
smt2_convt::identifiert
Definition: smt2_conv.h:223
boolbv_widtht::membert::width
std::size_t width
Definition: boolbv_width.h:41
to_address_of_expr
const address_of_exprt & to_address_of_expr(const exprt &expr)
Cast an exprt to an address_of_exprt.
Definition: pointer_expr.h:408
smt2_convt::find_symbols
void find_symbols(const exprt &expr)
Definition: smt2_conv.cpp:4800
extractbits_exprt
Extracts a sub-range of a bit-vector operand.
Definition: bitvector_expr.h:447
exprt::operands
operandst & operands()
Definition: expr.h:94
smt2_convt::solvert
solvert
Definition: smt2_conv.h:39
pointer_logict::add_object
mp_integer add_object(const exprt &expr)
Definition: pointer_logic.cpp:44
float_bv.h
let_exprt::values
operandst & values()
Definition: std_expr.h:3212
to_named_term_expr
const named_term_exprt & to_named_term_expr(const exprt &expr)
Cast an exprt to a named_term_exprt.
Definition: std_expr.h:3602
is_null_pointer
bool is_null_pointer(const constant_exprt &expr)
Returns true if expr has a pointer type and a value NULL; it also returns true when expr has value ze...
Definition: expr_util.cpp:315
integer2binary
const std::string integer2binary(const mp_integer &n, std::size_t width)
Definition: mp_arith.cpp:64
index_exprt
Array index operator.
Definition: std_expr.h:1409
smt2_convt::boolean_assignment
std::vector< bool > boolean_assignment
Definition: smt2_conv.h:267
address_of_exprt
Operator to return the address of an object.
Definition: pointer_expr.h:370
s2
int16_t s2
Definition: bytecode_info.h:60
smt2_convt::use_as_const
bool use_as_const
Definition: smt2_conv.h:63
typecast_exprt
Semantic type conversion.
Definition: std_expr.h:2016
pointer_typet
The pointer type These are both 'bitvector_typet' (they have a width) and 'type_with_subtypet' (they ...
Definition: pointer_expr.h:23
smt2_convt::object_sizes
std::map< object_size_exprt, irep_idt > object_sizes
Definition: smt2_conv.h:260
ieee_floatt::NaN
static ieee_floatt NaN(const ieee_float_spect &_spec)
Definition: ieee_float.h:209
true_exprt
The Boolean constant true.
Definition: std_expr.h:3007
constant_exprt
A constant literal expression.
Definition: std_expr.h:2941
is_smt2_simple_identifier
static bool is_smt2_simple_identifier(const std::string &identifier)
Definition: smt2_conv.cpp:909
is_true
bool is_true(const literalt &l)
Definition: literal.h:198
multi_ary_exprt::op0
exprt & op0()
Definition: std_expr.h:877
binary_exprt::op1
exprt & op1()
Definition: expr.h:128
to_isnan_expr
const isnan_exprt & to_isnan_expr(const exprt &expr)
Cast an exprt to a isnan_exprt.
Definition: floatbv_expr.h:113
to_multi_ary_expr
const multi_ary_exprt & to_multi_ary_expr(const exprt &expr)
Cast an exprt to a multi_ary_exprt.
Definition: std_expr.h:932
let_exprt::where
exprt & where()
convenience accessor for binding().where()
Definition: std_expr.h:3235
std_expr.h
to_binary_relation_expr
const binary_relation_exprt & to_binary_relation_expr(const exprt &expr)
Cast an exprt to a binary_relation_exprt.
Definition: std_expr.h:840
constant_exprt::set_value
void set_value(const irep_idt &value)
Definition: std_expr.h:2955
smt2_convt::unflatten
void unflatten(wheret, const typet &, unsigned nesting=0)
Definition: smt2_conv.cpp:4505
constant_exprt::get_value
const irep_idt & get_value() const
Definition: std_expr.h:2950
array_typet::index_type
typet index_type() const
The type of the index expressions into any instance of this type.
Definition: std_types.cpp:33
with_exprt::where
exprt & where()
Definition: std_expr.h:2444
member_offset
optionalt< mp_integer > member_offset(const struct_typet &type, const irep_idt &member, const namespacet &ns)
Definition: pointer_offset_size.cpp:24
boolbv_widtht::get_member
const membert & get_member(const struct_typet &type, const irep_idt &member) const
Definition: boolbv_width.cpp:230
array_exprt
Array constructor from list of elements.
Definition: std_expr.h:1562
smt2_convt::parse_union
exprt parse_union(const irept &s, const union_typet &type)
Definition: smt2_conv.cpp:561
make_range
ranget< iteratort > make_range(iteratort begin, iteratort end)
Definition: range.h:524
ieee_float_op_exprt::rhs
exprt & rhs()
Definition: floatbv_expr.h:385
smt2_convt::convert_floatbv_minus
void convert_floatbv_minus(const ieee_float_op_exprt &expr)
Definition: smt2_conv.cpp:3811
c_types.h
to_function_application_expr
const function_application_exprt & to_function_application_expr(const exprt &expr)
Cast an exprt to a function_application_exprt.
Definition: mathematical_expr.h:247
smt2_convt::flatten2bv
void flatten2bv(const exprt &)
Definition: smt2_conv.cpp:4416
isnormal_exprt
Evaluates to true if the operand is a normal number.
Definition: floatbv_expr.h:219
smt2_convt::smt2_identifiers
smt2_identifierst smt2_identifiers
Definition: smt2_conv.h:263
ieee_floatt::minus_infinity
static ieee_floatt minus_infinity(const ieee_float_spect &_spec)
Definition: ieee_float.h:215
let_exprt
A let expression.
Definition: std_expr.h:3141
pointer_logict::pointert
Definition: pointer_logic.h:28
bitwise_or
mp_integer bitwise_or(const mp_integer &a, const mp_integer &b)
bitwise 'or' of two nonnegative integers
Definition: mp_arith.cpp:215
literal_expr.h
to_count_leading_zeros_expr
const count_leading_zeros_exprt & to_count_leading_zeros_expr(const exprt &expr)
Cast an exprt to a count_leading_zeros_exprt.
Definition: bitvector_expr.h:1043
bitvector_expr.h
smt2_convt::convert_floatbv_typecast
void convert_floatbv_typecast(const floatbv_typecast_exprt &expr)
Definition: smt2_conv.cpp:2876
to_replication_expr
const replication_exprt & to_replication_expr(const exprt &expr)
Cast an exprt to a replication_exprt.
Definition: bitvector_expr.h:585
binary_exprt::op0
exprt & op0()
Definition: expr.h:125
to_struct_expr
const struct_exprt & to_struct_expr(const exprt &expr)
Cast an exprt to a struct_exprt.
Definition: std_expr.h:1842
validation_modet::INVARIANT
@ INVARIANT
to_abs_expr
const abs_exprt & to_abs_expr(const exprt &expr)
Cast an exprt to a abs_exprt.
Definition: std_expr.h:403
mod_exprt
Modulo defined as lhs-(rhs * truncate(lhs/rhs)).
Definition: std_expr.h:1167
smt2_convt::wheret::BEGIN
@ BEGIN
ieee_float_spect::f
std::size_t f
Definition: ieee_float.h:27
SMT2_TODO
#define SMT2_TODO(S)
Definition: smt2_conv.cpp:53
array_typet::element_type
const typet & element_type() const
The type of the elements of the array.
Definition: std_types.h:785
pointer_logict::pointer_expr
exprt pointer_expr(const pointert &pointer, const pointer_typet &type) const
Convert an (object,offset) pair to an expression.
Definition: pointer_logic.cpp:67
can_cast_expr< minus_overflow_exprt >
bool can_cast_expr< minus_overflow_exprt >(const exprt &base)
Definition: bitvector_expr.h:830
isnan_exprt
Evaluates to true if the operand is NaN.
Definition: floatbv_expr.h:87
smt2_convt::decision_procedure_text
std::string decision_procedure_text() const override
Return a textual description of the decision procedure.
Definition: smt2_conv.cpp:126
mathematical_expr.h
not_exprt
Boolean negation.
Definition: std_expr.h:2277
ieee_floatt::is_NaN
bool is_NaN() const
Definition: ieee_float.h:249
to_constant_expr
const constant_exprt & to_constant_expr(const exprt &expr)
Cast an exprt to a constant_exprt.
Definition: std_expr.h:2992
ieee_floatt::build
void build(const mp_integer &exp, const mp_integer &frac)
Definition: ieee_float.cpp:472
smt2_convt::convert_member
void convert_member(const member_exprt &expr)
Definition: smt2_conv.cpp:4355
smt2_convt::defined_expressions
defined_expressionst defined_expressions
Definition: smt2_conv.h:254