CBMC
goto_check_c.cpp
Go to the documentation of this file.
1 /*******************************************************************\
2 
3 Module: Checks for Errors in C/C++ Programs
4 
5 Author: Daniel Kroening, kroening@kroening.com
6 
7 \*******************************************************************/
8 
11 
12 #include "goto_check_c.h"
13 
14 #include <util/arith_tools.h>
15 #include <util/array_name.h>
16 #include <util/bitvector_expr.h>
17 #include <util/c_types.h>
18 #include <util/config.h>
19 #include <util/cprover_prefix.h>
20 #include <util/expr_util.h>
21 #include <util/find_symbols.h>
22 #include <util/floatbv_expr.h>
23 #include <util/ieee_float.h>
24 #include <util/invariant.h>
25 #include <util/make_unique.h>
26 #include <util/mathematical_expr.h>
27 #include <util/message.h>
28 #include <util/options.h>
29 #include <util/pointer_expr.h>
32 #include <util/prefix.h>
33 #include <util/simplify_expr.h>
34 #include <util/std_code.h>
35 #include <util/std_expr.h>
36 
39 
41 #include <langapi/language.h>
42 #include <langapi/mode.h>
43 
44 #include <algorithm>
45 #include <optional>
46 
48 {
49 public:
51  const namespacet &_ns,
52  const optionst &_options,
53  message_handlert &_message_handler)
54  : ns(_ns), local_bitvector_analysis(nullptr), log(_message_handler)
55  {
56  enable_bounds_check = _options.get_bool_option("bounds-check");
57  enable_pointer_check = _options.get_bool_option("pointer-check");
58  enable_memory_leak_check = _options.get_bool_option("memory-leak-check");
59  enable_div_by_zero_check = _options.get_bool_option("div-by-zero-check");
60  enable_enum_range_check = _options.get_bool_option("enum-range-check");
62  _options.get_bool_option("signed-overflow-check");
64  _options.get_bool_option("unsigned-overflow-check");
66  _options.get_bool_option("pointer-overflow-check");
67  enable_conversion_check = _options.get_bool_option("conversion-check");
69  _options.get_bool_option("undefined-shift-check");
71  _options.get_bool_option("float-overflow-check");
72  enable_simplify = _options.get_bool_option("simplify");
73  enable_nan_check = _options.get_bool_option("nan-check");
74  retain_trivial = _options.get_bool_option("retain-trivial-checks");
75  enable_assert_to_assume = _options.get_bool_option("assert-to-assume");
76  error_labels = _options.get_list_option("error-label");
78  _options.get_bool_option("pointer-primitive-check");
79  }
80 
82 
83  void goto_check(
84  const irep_idt &function_identifier,
85  goto_functiont &goto_function);
86 
92  void collect_allocations(const goto_functionst &goto_functions);
93 
94 protected:
95  const namespacet &ns;
96  std::unique_ptr<local_bitvector_analysist> local_bitvector_analysis;
99 
100  using guardt = std::function<exprt(exprt)>;
101  const guardt identity = [](exprt expr) { return expr; };
102 
108  void check_rec_address(const exprt &expr, const guardt &guard);
109 
117  void check_rec_logical_op(const exprt &expr, const guardt &guard);
118 
125  void check_rec_if(const if_exprt &if_expr, const guardt &guard);
126 
137  bool check_rec_member(const member_exprt &member, const guardt &guard);
138 
143  void check_rec_div(const div_exprt &div_expr, const guardt &guard);
144 
149  void check_rec_arithmetic_op(const exprt &expr, const guardt &guard);
150 
156  void check_rec(const exprt &expr, const guardt &guard);
157 
160  void check(const exprt &expr);
161 
162  struct conditiont
163  {
164  conditiont(const exprt &_assertion, const std::string &_description)
165  : assertion(_assertion), description(_description)
166  {
167  }
168 
170  std::string description;
171  };
172 
173  using conditionst = std::list<conditiont>;
174 
175  void bounds_check(const exprt &, const guardt &);
176  void bounds_check_index(const index_exprt &, const guardt &);
177  void bounds_check_bit_count(const unary_exprt &, const guardt &);
178  void div_by_zero_check(const div_exprt &, const guardt &);
179  void mod_by_zero_check(const mod_exprt &, const guardt &);
180  void mod_overflow_check(const mod_exprt &, const guardt &);
181  void enum_range_check(const exprt &, const guardt &);
182  void undefined_shift_check(const shift_exprt &, const guardt &);
183  void pointer_rel_check(const binary_exprt &, const guardt &);
184  void pointer_overflow_check(const exprt &, const guardt &);
185 
192  const dereference_exprt &expr,
193  const exprt &src_expr,
194  const guardt &guard);
195 
201  void pointer_primitive_check(const exprt &expr, const guardt &guard);
202 
209  bool requires_pointer_primitive_check(const exprt &expr);
210 
212  get_pointer_is_null_condition(const exprt &address, const exprt &size);
214  const exprt &address,
215  const exprt &size);
217  const exprt &pointer,
218  const exprt &size);
219 
221  const exprt &address,
222  const exprt &size);
223  void integer_overflow_check(const exprt &, const guardt &);
224  void conversion_check(const exprt &, const guardt &);
225  void float_overflow_check(const exprt &, const guardt &);
226  void nan_check(const exprt &, const guardt &);
228 
229  std::string array_name(const exprt &);
230 
240  const exprt &asserted_expr,
241  const std::string &comment,
242  const std::string &property_class,
243  const source_locationt &source_location,
244  const exprt &src_expr,
245  const guardt &guard);
246 
248  typedef std::set<std::pair<exprt, exprt>> assertionst;
250 
255  void invalidate(const exprt &lhs);
256 
273 
275  std::map<irep_idt, bool *> name_to_flag{
276  {"bounds-check", &enable_bounds_check},
277  {"pointer-check", &enable_pointer_check},
278  {"memory-leak-check", &enable_memory_leak_check},
279  {"div-by-zero-check", &enable_div_by_zero_check},
280  {"enum-range-check", &enable_enum_range_check},
281  {"signed-overflow-check", &enable_signed_overflow_check},
282  {"unsigned-overflow-check", &enable_unsigned_overflow_check},
283  {"pointer-overflow-check", &enable_pointer_overflow_check},
284  {"conversion-check", &enable_conversion_check},
285  {"undefined-shift-check", &enable_undefined_shift_check},
286  {"float-overflow-check", &enable_float_overflow_check},
287  {"nan-check", &enable_nan_check},
288  {"pointer-primitive-check", &enable_pointer_primitive_check}};
289 
292 
293  // the first element of the pair is the base address,
294  // and the second is the size of the region
295  typedef std::pair<exprt, exprt> allocationt;
296  typedef std::list<allocationt> allocationst;
298 
300 
303  void add_active_named_check_pragmas(source_locationt &source_location) const;
304 
307  void
309 
311  typedef enum
312  {
316  } check_statust;
317 
320 
329  named_check_statust match_named_check(const irep_idt &named_check) const;
330 };
331 
340 {
341 public:
344  {
345  }
346 
352  void set_flag(bool &flag, bool new_value, const irep_idt &flag_name)
353  {
354  // make this a no-op if the flag is disabled
355  if(disabled_flags.find(&flag) != disabled_flags.end())
356  return;
357 
358  // detect double sets
359  INVARIANT(
360  flags_to_reset.find(&flag) == flags_to_reset.end(),
361  "Flag " + id2string(flag_name) + " set twice at \n" +
363  if(flag != new_value)
364  {
365  flags_to_reset[&flag] = flag;
366  flag = new_value;
367  }
368  }
369 
374  void disable_flag(bool &flag, const irep_idt &flag_name)
375  {
376  INVARIANT(
377  disabled_flags.find(&flag) == disabled_flags.end(),
378  "Flag " + id2string(flag_name) + " disabled twice at \n" +
380 
381  disabled_flags.insert(&flag);
382 
383  // If the flag has not already been set,
384  // we store its current value in the reset map.
385  // Otherwise, the reset map already holds
386  // the initial value we want to reset it to, keep it as is.
387  if(flags_to_reset.find(&flag) == flags_to_reset.end())
388  flags_to_reset[&flag] = flag;
389 
390  // set the flag to false in all cases.
391  flag = false;
392  }
393 
397  {
398  for(const auto &flag_pair : flags_to_reset)
399  *flag_pair.first = flag_pair.second;
400  }
401 
402 private:
404  std::map<bool *, bool> flags_to_reset;
405  std::set<bool *> disabled_flags;
406 };
407 
408 static exprt implication(exprt lhs, exprt rhs)
409 {
410  // rewrite a => (b => c) to (a && b) => c
411  if(rhs.id() == ID_implies)
412  {
413  const auto &rhs_implication = to_implies_expr(rhs);
414  return implies_exprt(
415  and_exprt(std::move(lhs), rhs_implication.lhs()), rhs_implication.rhs());
416  }
417  else
418  {
419  return implies_exprt(std::move(lhs), std::move(rhs));
420  }
421 }
422 
424 {
425  for(const auto &gf_entry : goto_functions.function_map)
426  {
427  for(const auto &instruction : gf_entry.second.body.instructions)
428  {
429  if(!instruction.is_function_call())
430  continue;
431 
432  const auto &function = instruction.call_function();
433  if(
434  function.id() != ID_symbol ||
436  "allocated_memory")
437  continue;
438 
439  const auto function_line = function.source_location().as_string();
440  log.warning() << "**** WARNING: `" CPROVER_PREFIX "allocated_memory' in "
441  << function_line << messaget::eom;
442  log.warning() << "**** WARNING: `" CPROVER_PREFIX
443  "allocated_memory' is "
444  "deprecated and scheduled for deletion "
445  << "in version 6 and upwards." << messaget::eom;
446  log.warning() << "Please avoid using this intrinsic. For more "
447  "information, please check issue "
448  << "cbmc#6872 in Github" << messaget::eom;
449 
450  const code_function_callt::argumentst &args =
451  instruction.call_arguments();
452  if(
453  args.size() != 2 || args[0].type().id() != ID_unsignedbv ||
454  args[1].type().id() != ID_unsignedbv)
455  throw "expected two unsigned arguments to " CPROVER_PREFIX
456  "allocated_memory";
457 
459  args[0].type() == args[1].type(),
460  "arguments of allocated_memory must have same type");
461  allocations.push_back({args[0], args[1]});
462  }
463  }
464 }
465 
467 {
468  if(lhs.id() == ID_index)
469  invalidate(to_index_expr(lhs).array());
470  else if(lhs.id() == ID_member)
471  invalidate(to_member_expr(lhs).struct_op());
472  else if(lhs.id() == ID_symbol)
473  {
474  // clear all assertions about 'symbol'
475  const irep_idt &lhs_id = to_symbol_expr(lhs).get_identifier();
476 
477  for(auto it = assertions.begin(); it != assertions.end();)
478  {
479  if(
480  has_symbol_expr(it->second, lhs_id, false) ||
481  has_subexpr(it->second, ID_dereference))
482  {
483  it = assertions.erase(it);
484  }
485  else
486  ++it;
487  }
488  }
489  else
490  {
491  // give up, clear all
492  assertions.clear();
493  }
494 }
495 
497  const div_exprt &expr,
498  const guardt &guard)
499 {
501  return;
502 
503  // add divison by zero subgoal
504 
505  exprt zero = from_integer(0, expr.op1().type());
506  const notequal_exprt inequality(expr.op1(), std::move(zero));
507 
509  inequality,
510  "division by zero",
511  "division-by-zero",
512  expr.find_source_location(),
513  expr,
514  guard);
515 }
516 
517 void goto_check_ct::enum_range_check(const exprt &expr, const guardt &guard)
518 {
520  return;
521 
522  const c_enum_tag_typet &c_enum_tag_type = to_c_enum_tag_type(expr.type());
523  const c_enum_typet &c_enum_type = ns.follow_tag(c_enum_tag_type);
524 
525  const c_enum_typet::memberst enum_values = c_enum_type.members();
526 
527  std::vector<exprt> disjuncts;
528  for(const auto &enum_value : enum_values)
529  {
530  const constant_exprt val{enum_value.get_value(), c_enum_tag_type};
531  disjuncts.push_back(equal_exprt(expr, val));
532  }
533 
534  const exprt check = disjunction(disjuncts);
535 
537  check,
538  "enum range check",
539  "enum-range-check",
540  expr.find_source_location(),
541  expr,
542  guard);
543 }
544 
546  const shift_exprt &expr,
547  const guardt &guard)
548 {
550  return;
551 
552  // Undefined for all types and shifts if distance exceeds width,
553  // and also undefined for negative distances.
554 
555  const typet &distance_type = expr.distance().type();
556 
557  if(distance_type.id() == ID_signedbv)
558  {
559  binary_relation_exprt inequality(
560  expr.distance(), ID_ge, from_integer(0, distance_type));
561 
563  inequality,
564  "shift distance is negative",
565  "undefined-shift",
566  expr.find_source_location(),
567  expr,
568  guard);
569  }
570 
571  const typet &op_type = expr.op().type();
572 
573  if(op_type.id() == ID_unsignedbv || op_type.id() == ID_signedbv)
574  {
575  exprt width_expr =
576  from_integer(to_bitvector_type(op_type).get_width(), distance_type);
577 
579  binary_relation_exprt(expr.distance(), ID_lt, std::move(width_expr)),
580  "shift distance too large",
581  "undefined-shift",
582  expr.find_source_location(),
583  expr,
584  guard);
585 
586  if(op_type.id() == ID_signedbv && expr.id() == ID_shl)
587  {
588  binary_relation_exprt inequality(
589  expr.op(), ID_ge, from_integer(0, op_type));
590 
592  inequality,
593  "shift operand is negative",
594  "undefined-shift",
595  expr.find_source_location(),
596  expr,
597  guard);
598  }
599  }
600  else
601  {
603  false_exprt(),
604  "shift of non-integer type",
605  "undefined-shift",
606  expr.find_source_location(),
607  expr,
608  guard);
609  }
610 }
611 
613  const mod_exprt &expr,
614  const guardt &guard)
615 {
617  return;
618 
619  // add divison by zero subgoal
620 
621  exprt zero = from_integer(0, expr.divisor().type());
622  const notequal_exprt inequality(expr.divisor(), std::move(zero));
623 
625  inequality,
626  "division by zero",
627  "division-by-zero",
628  expr.find_source_location(),
629  expr,
630  guard);
631 }
632 
635  const mod_exprt &expr,
636  const guardt &guard)
637 {
639  return;
640 
641  const auto &type = expr.type();
642 
643  if(type.id() == ID_signedbv)
644  {
645  // INT_MIN % -1 is, in principle, defined to be zero in
646  // ANSI C, C99, C++98, and C++11. Most compilers, however,
647  // fail to produce 0, and in some cases generate an exception.
648  // C11 explicitly makes this case undefined.
649 
650  notequal_exprt int_min_neq(
651  expr.op0(), to_signedbv_type(type).smallest_expr());
652 
653  notequal_exprt minus_one_neq(
654  expr.op1(), from_integer(-1, expr.op1().type()));
655 
657  or_exprt(int_min_neq, minus_one_neq),
658  "result of signed mod is not representable",
659  "overflow",
660  expr.find_source_location(),
661  expr,
662  guard);
663  }
664 }
665 
666 void goto_check_ct::conversion_check(const exprt &expr, const guardt &guard)
667 {
669  return;
670 
671  // First, check type.
672  const typet &type = expr.type();
673 
674  if(type.id() != ID_signedbv && type.id() != ID_unsignedbv)
675  return;
676 
677  if(expr.id() == ID_typecast)
678  {
679  const auto &op = to_typecast_expr(expr).op();
680 
681  // conversion to signed int may overflow
682  const typet &old_type = op.type();
683 
684  if(type.id() == ID_signedbv)
685  {
686  std::size_t new_width = to_signedbv_type(type).get_width();
687 
688  if(old_type.id() == ID_signedbv) // signed -> signed
689  {
690  std::size_t old_width = to_signedbv_type(old_type).get_width();
691  if(new_width >= old_width)
692  return; // always ok
693 
694  const binary_relation_exprt no_overflow_upper(
695  op, ID_le, from_integer(power(2, new_width - 1) - 1, old_type));
696 
697  const binary_relation_exprt no_overflow_lower(
698  op, ID_ge, from_integer(-power(2, new_width - 1), old_type));
699 
701  and_exprt(no_overflow_lower, no_overflow_upper),
702  "arithmetic overflow on signed type conversion",
703  "overflow",
704  expr.find_source_location(),
705  expr,
706  guard);
707  }
708  else if(old_type.id() == ID_unsignedbv) // unsigned -> signed
709  {
710  std::size_t old_width = to_unsignedbv_type(old_type).get_width();
711  if(new_width >= old_width + 1)
712  return; // always ok
713 
714  const binary_relation_exprt no_overflow_upper(
715  op, ID_le, from_integer(power(2, new_width - 1) - 1, old_type));
716 
718  no_overflow_upper,
719  "arithmetic overflow on unsigned to signed type conversion",
720  "overflow",
721  expr.find_source_location(),
722  expr,
723  guard);
724  }
725  else if(old_type.id() == ID_floatbv) // float -> signed
726  {
727  // Note that the fractional part is truncated!
728  ieee_floatt upper(to_floatbv_type(old_type));
729  upper.from_integer(power(2, new_width - 1));
730  const binary_relation_exprt no_overflow_upper(
731  op, ID_lt, upper.to_expr());
732 
733  ieee_floatt lower(to_floatbv_type(old_type));
734  lower.from_integer(-power(2, new_width - 1) - 1);
735  const binary_relation_exprt no_overflow_lower(
736  op, ID_gt, lower.to_expr());
737 
739  and_exprt(no_overflow_lower, no_overflow_upper),
740  "arithmetic overflow on float to signed integer type conversion",
741  "overflow",
742  expr.find_source_location(),
743  expr,
744  guard);
745  }
746  }
747  else if(type.id() == ID_unsignedbv)
748  {
749  std::size_t new_width = to_unsignedbv_type(type).get_width();
750 
751  if(old_type.id() == ID_signedbv) // signed -> unsigned
752  {
753  std::size_t old_width = to_signedbv_type(old_type).get_width();
754 
755  if(new_width >= old_width - 1)
756  {
757  // only need lower bound check
758  const binary_relation_exprt no_overflow_lower(
759  op, ID_ge, from_integer(0, old_type));
760 
762  no_overflow_lower,
763  "arithmetic overflow on signed to unsigned type conversion",
764  "overflow",
765  expr.find_source_location(),
766  expr,
767  guard);
768  }
769  else
770  {
771  // need both
772  const binary_relation_exprt no_overflow_upper(
773  op, ID_le, from_integer(power(2, new_width) - 1, old_type));
774 
775  const binary_relation_exprt no_overflow_lower(
776  op, ID_ge, from_integer(0, old_type));
777 
779  and_exprt(no_overflow_lower, no_overflow_upper),
780  "arithmetic overflow on signed to unsigned type conversion",
781  "overflow",
782  expr.find_source_location(),
783  expr,
784  guard);
785  }
786  }
787  else if(old_type.id() == ID_unsignedbv) // unsigned -> unsigned
788  {
789  std::size_t old_width = to_unsignedbv_type(old_type).get_width();
790  if(new_width >= old_width)
791  return; // always ok
792 
793  const binary_relation_exprt no_overflow_upper(
794  op, ID_le, from_integer(power(2, new_width) - 1, old_type));
795 
797  no_overflow_upper,
798  "arithmetic overflow on unsigned to unsigned type conversion",
799  "overflow",
800  expr.find_source_location(),
801  expr,
802  guard);
803  }
804  else if(old_type.id() == ID_floatbv) // float -> unsigned
805  {
806  // Note that the fractional part is truncated!
807  ieee_floatt upper(to_floatbv_type(old_type));
808  upper.from_integer(power(2, new_width) - 1);
809  const binary_relation_exprt no_overflow_upper(
810  op, ID_lt, upper.to_expr());
811 
812  ieee_floatt lower(to_floatbv_type(old_type));
813  lower.from_integer(-1);
814  const binary_relation_exprt no_overflow_lower(
815  op, ID_gt, lower.to_expr());
816 
818  and_exprt(no_overflow_lower, no_overflow_upper),
819  "arithmetic overflow on float to unsigned integer type conversion",
820  "overflow",
821  expr.find_source_location(),
822  expr,
823  guard);
824  }
825  }
826  }
827 }
828 
830  const exprt &expr,
831  const guardt &guard)
832 {
834  return;
835 
836  // First, check type.
837  const typet &type = expr.type();
838 
839  if(type.id() == ID_signedbv && !enable_signed_overflow_check)
840  return;
841 
842  if(type.id() == ID_unsignedbv && !enable_unsigned_overflow_check)
843  return;
844 
845  // add overflow subgoal
846 
847  if(expr.id() == ID_div)
848  {
849  // undefined for signed division INT_MIN/-1
850  if(type.id() == ID_signedbv)
851  {
852  const auto &div_expr = to_div_expr(expr);
853 
854  equal_exprt int_min_eq(
855  div_expr.dividend(), to_signedbv_type(type).smallest_expr());
856 
857  equal_exprt minus_one_eq(div_expr.divisor(), from_integer(-1, type));
858 
860  not_exprt(and_exprt(int_min_eq, minus_one_eq)),
861  "arithmetic overflow on signed division",
862  "overflow",
863  expr.find_source_location(),
864  expr,
865  guard);
866  }
867 
868  return;
869  }
870  else if(expr.id() == ID_unary_minus)
871  {
872  if(type.id() == ID_signedbv)
873  {
874  // overflow on unary- on signed integers can only happen with the
875  // smallest representable number 100....0
876 
877  equal_exprt int_min_eq(
878  to_unary_minus_expr(expr).op(), to_signedbv_type(type).smallest_expr());
879 
881  not_exprt(int_min_eq),
882  "arithmetic overflow on signed unary minus",
883  "overflow",
884  expr.find_source_location(),
885  expr,
886  guard);
887  }
888  else if(type.id() == ID_unsignedbv)
889  {
890  // Overflow on unary- on unsigned integers happens for all operands
891  // that are not zero.
892 
893  notequal_exprt not_eq_zero(
894  to_unary_minus_expr(expr).op(), to_unsignedbv_type(type).zero_expr());
895 
897  not_eq_zero,
898  "arithmetic overflow on unsigned unary minus",
899  "overflow",
900  expr.find_source_location(),
901  expr,
902  guard);
903  }
904 
905  return;
906  }
907  else if(expr.id() == ID_shl)
908  {
909  if(type.id() == ID_signedbv)
910  {
911  const auto &shl_expr = to_shl_expr(expr);
912  const auto &op = shl_expr.op();
913  const auto &op_type = to_signedbv_type(type);
914  const auto op_width = op_type.get_width();
915  const auto &distance = shl_expr.distance();
916  const auto &distance_type = distance.type();
917 
918  // a left shift of a negative value is undefined;
919  // yet this isn't an overflow
920  exprt neg_value_shift;
921 
922  if(op_type.id() == ID_unsignedbv)
923  neg_value_shift = false_exprt();
924  else
925  neg_value_shift =
926  binary_relation_exprt(op, ID_lt, from_integer(0, op_type));
927 
928  // a shift with negative distance is undefined;
929  // yet this isn't an overflow
930  exprt neg_dist_shift;
931 
932  if(distance_type.id() == ID_unsignedbv)
933  neg_dist_shift = false_exprt();
934  else
935  {
936  neg_dist_shift = binary_relation_exprt(
937  distance, ID_lt, from_integer(0, distance_type));
938  }
939 
940  // shifting a non-zero value by more than its width is undefined;
941  // yet this isn't an overflow
942  const exprt dist_too_large = binary_predicate_exprt(
943  distance, ID_gt, from_integer(op_width, distance_type));
944 
945  const exprt op_zero = equal_exprt(op, from_integer(0, op_type));
946 
947  auto new_type = to_bitvector_type(op_type);
948  new_type.set_width(op_width * 2);
949 
950  const exprt op_ext = typecast_exprt(op, new_type);
951 
952  const exprt op_ext_shifted = shl_exprt(op_ext, distance);
953 
954  // The semantics of signed left shifts are contentious for the case
955  // that a '1' is shifted into the signed bit.
956  // Assuming 32-bit integers, 1<<31 is implementation-defined
957  // in ANSI C and C++98, but is explicitly undefined by C99,
958  // C11 and C++11.
959  bool allow_shift_into_sign_bit = true;
960 
961  if(mode == ID_C)
962  {
963  if(
966  {
967  allow_shift_into_sign_bit = false;
968  }
969  }
970  else if(mode == ID_cpp)
971  {
972  if(
976  {
977  allow_shift_into_sign_bit = false;
978  }
979  }
980 
981  const unsigned number_of_top_bits =
982  allow_shift_into_sign_bit ? op_width : op_width + 1;
983 
984  const exprt top_bits = extractbits_exprt(
985  op_ext_shifted,
986  new_type.get_width() - 1,
987  new_type.get_width() - number_of_top_bits,
988  unsignedbv_typet(number_of_top_bits));
989 
990  const exprt top_bits_zero =
991  equal_exprt(top_bits, from_integer(0, top_bits.type()));
992 
993  // a negative distance shift isn't an overflow;
994  // a negative value shift isn't an overflow;
995  // a shift that's too far isn't an overflow;
996  // a shift of zero isn't overflow;
997  // else check the top bits
999  disjunction(
1000  {neg_value_shift,
1001  neg_dist_shift,
1002  dist_too_large,
1003  op_zero,
1004  top_bits_zero}),
1005  "arithmetic overflow on signed shl",
1006  "overflow",
1007  expr.find_source_location(),
1008  expr,
1009  guard);
1010  }
1011 
1012  return;
1013  }
1014 
1015  multi_ary_exprt overflow(
1016  "overflow-" + expr.id_string(), expr.operands(), bool_typet());
1017 
1018  if(expr.operands().size() >= 3)
1019  {
1020  // The overflow checks are binary!
1021  // We break these up.
1022 
1023  for(std::size_t i = 1; i < expr.operands().size(); i++)
1024  {
1025  exprt tmp;
1026 
1027  if(i == 1)
1028  tmp = to_multi_ary_expr(expr).op0();
1029  else
1030  {
1031  tmp = expr;
1032  tmp.operands().resize(i);
1033  }
1034 
1035  std::string kind = type.id() == ID_unsignedbv ? "unsigned" : "signed";
1036 
1038  not_exprt{binary_overflow_exprt{tmp, expr.id(), expr.operands()[i]}},
1039  "arithmetic overflow on " + kind + " " + expr.id_string(),
1040  "overflow",
1041  expr.find_source_location(),
1042  expr,
1043  guard);
1044  }
1045  }
1046  else if(expr.operands().size() == 2)
1047  {
1048  std::string kind = type.id() == ID_unsignedbv ? "unsigned" : "signed";
1049 
1050  const binary_exprt &bexpr = to_binary_expr(expr);
1052  not_exprt{binary_overflow_exprt{bexpr.lhs(), expr.id(), bexpr.rhs()}},
1053  "arithmetic overflow on " + kind + " " + expr.id_string(),
1054  "overflow",
1055  expr.find_source_location(),
1056  expr,
1057  guard);
1058  }
1059  else
1060  {
1061  PRECONDITION(expr.id() == ID_unary_minus);
1062 
1063  std::string kind = type.id() == ID_unsignedbv ? "unsigned" : "signed";
1064 
1067  "arithmetic overflow on " + kind + " " + expr.id_string(),
1068  "overflow",
1069  expr.find_source_location(),
1070  expr,
1071  guard);
1072  }
1073 }
1074 
1075 void goto_check_ct::float_overflow_check(const exprt &expr, const guardt &guard)
1076 {
1078  return;
1079 
1080  // First, check type.
1081  const typet &type = expr.type();
1082 
1083  if(type.id() != ID_floatbv)
1084  return;
1085 
1086  // add overflow subgoal
1087 
1088  if(expr.id() == ID_typecast)
1089  {
1090  // Can overflow if casting from larger
1091  // to smaller type.
1092  const auto &op = to_typecast_expr(expr).op();
1093  if(op.type().id() == ID_floatbv)
1094  {
1095  // float-to-float
1096  or_exprt overflow_check{isinf_exprt(op), not_exprt(isinf_exprt(expr))};
1097 
1099  std::move(overflow_check),
1100  "arithmetic overflow on floating-point typecast",
1101  "overflow",
1102  expr.find_source_location(),
1103  expr,
1104  guard);
1105  }
1106  else
1107  {
1108  // non-float-to-float
1110  not_exprt(isinf_exprt(expr)),
1111  "arithmetic overflow on floating-point typecast",
1112  "overflow",
1113  expr.find_source_location(),
1114  expr,
1115  guard);
1116  }
1117 
1118  return;
1119  }
1120  else if(expr.id() == ID_div)
1121  {
1122  // Can overflow if dividing by something small
1123  or_exprt overflow_check(
1124  isinf_exprt(to_div_expr(expr).dividend()), not_exprt(isinf_exprt(expr)));
1125 
1127  std::move(overflow_check),
1128  "arithmetic overflow on floating-point division",
1129  "overflow",
1130  expr.find_source_location(),
1131  expr,
1132  guard);
1133 
1134  return;
1135  }
1136  else if(expr.id() == ID_mod)
1137  {
1138  // Can't overflow
1139  return;
1140  }
1141  else if(expr.id() == ID_unary_minus)
1142  {
1143  // Can't overflow
1144  return;
1145  }
1146  else if(expr.id() == ID_plus || expr.id() == ID_mult || expr.id() == ID_minus)
1147  {
1148  if(expr.operands().size() == 2)
1149  {
1150  // Can overflow
1151  or_exprt overflow_check(
1152  isinf_exprt(to_binary_expr(expr).op0()),
1153  isinf_exprt(to_binary_expr(expr).op1()),
1154  not_exprt(isinf_exprt(expr)));
1155 
1156  std::string kind = expr.id() == ID_plus
1157  ? "addition"
1158  : expr.id() == ID_minus
1159  ? "subtraction"
1160  : expr.id() == ID_mult ? "multiplication" : "";
1161 
1163  std::move(overflow_check),
1164  "arithmetic overflow on floating-point " + kind,
1165  "overflow",
1166  expr.find_source_location(),
1167  expr,
1168  guard);
1169 
1170  return;
1171  }
1172  else if(expr.operands().size() >= 3)
1173  {
1174  DATA_INVARIANT(expr.id() != ID_minus, "minus expression must be binary");
1175 
1176  // break up
1177  float_overflow_check(make_binary(expr), guard);
1178  return;
1179  }
1180  }
1181 }
1182 
1183 void goto_check_ct::nan_check(const exprt &expr, const guardt &guard)
1184 {
1185  if(!enable_nan_check)
1186  return;
1187 
1188  // first, check type
1189  if(expr.type().id() != ID_floatbv)
1190  return;
1191 
1192  if(
1193  expr.id() != ID_plus && expr.id() != ID_mult && expr.id() != ID_div &&
1194  expr.id() != ID_minus)
1195  return;
1196 
1197  // add NaN subgoal
1198 
1199  exprt isnan;
1200 
1201  if(expr.id() == ID_div)
1202  {
1203  const auto &div_expr = to_div_expr(expr);
1204 
1205  // there a two ways to get a new NaN on division:
1206  // 0/0 = NaN and x/inf = NaN
1207  // (note that x/0 = +-inf for x!=0 and x!=inf)
1208  const and_exprt zero_div_zero(
1210  div_expr.op0(), from_integer(0, div_expr.dividend().type())),
1212  div_expr.op1(), from_integer(0, div_expr.divisor().type())));
1213 
1214  const isinf_exprt div_inf(div_expr.op1());
1215 
1216  isnan = or_exprt(zero_div_zero, div_inf);
1217  }
1218  else if(expr.id() == ID_mult)
1219  {
1220  if(expr.operands().size() >= 3)
1221  return nan_check(make_binary(expr), guard);
1222 
1223  const auto &mult_expr = to_mult_expr(expr);
1224 
1225  // Inf * 0 is NaN
1226  const and_exprt inf_times_zero(
1227  isinf_exprt(mult_expr.op0()),
1229  mult_expr.op1(), from_integer(0, mult_expr.op1().type())));
1230 
1231  const and_exprt zero_times_inf(
1233  mult_expr.op1(), from_integer(0, mult_expr.op1().type())),
1234  isinf_exprt(mult_expr.op0()));
1235 
1236  isnan = or_exprt(inf_times_zero, zero_times_inf);
1237  }
1238  else if(expr.id() == ID_plus)
1239  {
1240  if(expr.operands().size() >= 3)
1241  return nan_check(make_binary(expr), guard);
1242 
1243  const auto &plus_expr = to_plus_expr(expr);
1244 
1245  // -inf + +inf = NaN and +inf + -inf = NaN,
1246  // i.e., signs differ
1247  ieee_float_spect spec = ieee_float_spect(to_floatbv_type(plus_expr.type()));
1248  exprt plus_inf = ieee_floatt::plus_infinity(spec).to_expr();
1249  exprt minus_inf = ieee_floatt::minus_infinity(spec).to_expr();
1250 
1251  isnan = or_exprt(
1252  and_exprt(
1253  equal_exprt(plus_expr.op0(), minus_inf),
1254  equal_exprt(plus_expr.op1(), plus_inf)),
1255  and_exprt(
1256  equal_exprt(plus_expr.op0(), plus_inf),
1257  equal_exprt(plus_expr.op1(), minus_inf)));
1258  }
1259  else if(expr.id() == ID_minus)
1260  {
1261  // +inf - +inf = NaN and -inf - -inf = NaN,
1262  // i.e., signs match
1263 
1264  const auto &minus_expr = to_minus_expr(expr);
1265 
1266  ieee_float_spect spec =
1267  ieee_float_spect(to_floatbv_type(minus_expr.type()));
1268  exprt plus_inf = ieee_floatt::plus_infinity(spec).to_expr();
1269  exprt minus_inf = ieee_floatt::minus_infinity(spec).to_expr();
1270 
1271  isnan = or_exprt(
1272  and_exprt(
1273  equal_exprt(minus_expr.lhs(), plus_inf),
1274  equal_exprt(minus_expr.rhs(), plus_inf)),
1275  and_exprt(
1276  equal_exprt(minus_expr.lhs(), minus_inf),
1277  equal_exprt(minus_expr.rhs(), minus_inf)));
1278  }
1279  else
1280  UNREACHABLE;
1281 
1283  boolean_negate(isnan),
1284  "NaN on " + expr.id_string(),
1285  "NaN",
1286  expr.find_source_location(),
1287  expr,
1288  guard);
1289 }
1290 
1292  const binary_exprt &expr,
1293  const guardt &guard)
1294 {
1296  return;
1297 
1298  if(
1299  expr.op0().type().id() == ID_pointer &&
1300  expr.op1().type().id() == ID_pointer)
1301  {
1302  // add same-object subgoal
1303 
1304  exprt same_object = ::same_object(expr.op0(), expr.op1());
1305 
1307  same_object,
1308  "same object violation",
1309  "pointer",
1310  expr.find_source_location(),
1311  expr,
1312  guard);
1313 
1314  for(const auto &pointer : expr.operands())
1315  {
1316  // just this particular byte must be within object bounds or one past the
1317  // end
1318  const auto size = from_integer(0, size_type());
1319  auto conditions = get_pointer_dereferenceable_conditions(pointer, size);
1320 
1321  for(const auto &c : conditions)
1322  {
1324  c.assertion,
1325  "pointer relation: " + c.description,
1326  "pointer arithmetic",
1327  expr.find_source_location(),
1328  pointer,
1329  guard);
1330  }
1331  }
1332  }
1333 }
1334 
1336  const exprt &expr,
1337  const guardt &guard)
1338 {
1340  return;
1341 
1342  if(expr.id() != ID_plus && expr.id() != ID_minus)
1343  return;
1344 
1346  expr.operands().size() == 2,
1347  "pointer arithmetic expected to have exactly 2 operands");
1348 
1349  // multiplying the offset by the object size must not result in arithmetic
1350  // overflow
1351  const typet &object_type = to_pointer_type(expr.type()).base_type();
1352  if(object_type.id() != ID_empty)
1353  {
1354  auto size_of_expr_opt = size_of_expr(object_type, ns);
1355  CHECK_RETURN(size_of_expr_opt.has_value());
1356  exprt object_size = size_of_expr_opt.value();
1357 
1358  const binary_exprt &binary_expr = to_binary_expr(expr);
1359  exprt offset_operand = binary_expr.lhs().type().id() == ID_pointer
1360  ? binary_expr.rhs()
1361  : binary_expr.lhs();
1362  mult_exprt mul{
1363  offset_operand,
1365  mul.add_source_location() = offset_operand.source_location();
1366 
1367  flag_overridet override_overflow(offset_operand.source_location());
1368  override_overflow.set_flag(
1369  enable_signed_overflow_check, true, "signed_overflow_check");
1370  override_overflow.set_flag(
1371  enable_unsigned_overflow_check, true, "unsigned_overflow_check");
1372  integer_overflow_check(mul, guard);
1373  }
1374 
1375  // the result must be within object bounds or one past the end
1376  const auto size = from_integer(0, size_type());
1377  auto conditions = get_pointer_dereferenceable_conditions(expr, size);
1378 
1379  for(const auto &c : conditions)
1380  {
1382  c.assertion,
1383  "pointer arithmetic: " + c.description,
1384  "pointer arithmetic",
1385  expr.find_source_location(),
1386  expr,
1387  guard);
1388  }
1389 }
1390 
1392  const dereference_exprt &expr,
1393  const exprt &src_expr,
1394  const guardt &guard)
1395 {
1397  return;
1398 
1399  const exprt &pointer = expr.pointer();
1400 
1401  exprt size;
1402 
1403  if(expr.type().id() == ID_empty)
1404  {
1405  // a dereference *p (with p being a pointer to void) is valid if p points to
1406  // valid memory (of any size). the smallest possible size of the memory
1407  // segment p could be pointing to is 1, hence we use this size for the
1408  // address check
1409  size = from_integer(1, size_type());
1410  }
1411  else
1412  {
1413  auto size_of_expr_opt = size_of_expr(expr.type(), ns);
1414  CHECK_RETURN(size_of_expr_opt.has_value());
1415  size = size_of_expr_opt.value();
1416  }
1417 
1418  auto conditions = get_pointer_dereferenceable_conditions(pointer, size);
1419 
1420  for(const auto &c : conditions)
1421  {
1423  c.assertion,
1424  "dereference failure: " + c.description,
1425  "pointer dereference",
1426  src_expr.find_source_location(),
1427  src_expr,
1428  guard);
1429  }
1430 }
1431 
1433  const exprt &expr,
1434  const guardt &guard)
1435 {
1437  return;
1438 
1439  if(expr.source_location().is_built_in())
1440  return;
1441 
1442  const exprt pointer =
1443  (expr.id() == ID_r_ok || expr.id() == ID_w_ok || expr.id() == ID_rw_ok)
1444  ? to_r_or_w_ok_expr(expr).pointer()
1445  : to_unary_expr(expr).op();
1446 
1447  CHECK_RETURN(pointer.type().id() == ID_pointer);
1448 
1449  if(pointer.id() == ID_symbol)
1450  {
1451  const auto &symbol_expr = to_symbol_expr(pointer);
1452 
1453  if(has_prefix(id2string(symbol_expr.get_identifier()), CPROVER_PREFIX))
1454  return;
1455  }
1456 
1457  const auto size_of_expr_opt =
1458  size_of_expr(to_pointer_type(pointer.type()).base_type(), ns);
1459 
1460  const exprt size = !size_of_expr_opt.has_value()
1461  ? from_integer(1, size_type())
1462  : size_of_expr_opt.value();
1463 
1464  const conditionst &conditions =
1466  for(const auto &c : conditions)
1467  {
1469  or_exprt{null_object(pointer), c.assertion},
1470  c.description,
1471  "pointer primitives",
1472  expr.source_location(),
1473  expr,
1474  guard);
1475  }
1476 }
1477 
1479 {
1480  // we don't need to include the __CPROVER_same_object primitive here as it
1481  // is replaced by lower level primitives in the special function handling
1482  // during typechecking (see c_typecheck_expr.cpp)
1483 
1484  // pointer_object and pointer_offset are well-defined for an arbitrary
1485  // pointer-typed operand (and the operands themselves have been checked
1486  // separately for, e.g., invalid pointer dereferencing via check_rec):
1487  // pointer_object and pointer_offset just extract a subset of bits from the
1488  // pointer. If that pointer was unconstrained (non-deterministic), the result
1489  // will equally be non-deterministic.
1490  return can_cast_expr<object_size_exprt>(expr) || expr.id() == ID_r_ok ||
1491  expr.id() == ID_w_ok || expr.id() == ID_rw_ok ||
1492  expr.id() == ID_is_dynamic_object;
1493 }
1494 
1497  const exprt &address,
1498  const exprt &size)
1499 {
1500  auto conditions =
1502  if(auto maybe_null_condition = get_pointer_is_null_condition(address, size))
1503  {
1504  conditions.push_front(*maybe_null_condition);
1505  }
1506  return conditions;
1507 }
1508 
1509 std::string goto_check_ct::array_name(const exprt &expr)
1510 {
1511  return ::array_name(ns, expr);
1512 }
1513 
1514 void goto_check_ct::bounds_check(const exprt &expr, const guardt &guard)
1515 {
1516  if(!enable_bounds_check)
1517  return;
1518 
1519  if(expr.id() == ID_index)
1520  bounds_check_index(to_index_expr(expr), guard);
1521  else if(
1522  (expr.id() == ID_count_leading_zeros &&
1523  !to_count_leading_zeros_expr(expr).zero_permitted()) ||
1524  (expr.id() == ID_count_trailing_zeros &&
1525  !to_count_trailing_zeros_expr(expr).zero_permitted()))
1526  {
1527  bounds_check_bit_count(to_unary_expr(expr), guard);
1528  }
1529 }
1530 
1532  const index_exprt &expr,
1533  const guardt &guard)
1534 {
1535  const typet &array_type = expr.array().type();
1536 
1537  if(array_type.id() == ID_pointer)
1538  throw "index got pointer as array type";
1539  else if(array_type.id() != ID_array && array_type.id() != ID_vector)
1540  throw "bounds check expected array or vector type, got " +
1541  array_type.id_string();
1542 
1543  std::string name = array_name(expr.array());
1544 
1545  const exprt &index = expr.index();
1547  ode.build(expr, ns);
1548 
1549  if(index.type().id() != ID_unsignedbv)
1550  {
1551  // we undo typecasts to signedbv
1552  if(
1553  index.id() == ID_typecast &&
1554  to_typecast_expr(index).op().type().id() == ID_unsignedbv)
1555  {
1556  // ok
1557  }
1558  else
1559  {
1560  const auto i = numeric_cast<mp_integer>(index);
1561 
1562  if(!i.has_value() || *i < 0)
1563  {
1564  exprt effective_offset = ode.offset();
1565 
1566  if(ode.root_object().id() == ID_dereference)
1567  {
1568  exprt p_offset =
1570 
1571  effective_offset = plus_exprt{
1572  p_offset,
1574  effective_offset, p_offset.type())};
1575  }
1576 
1577  exprt zero = from_integer(0, ode.offset().type());
1578 
1579  // the final offset must not be negative
1580  binary_relation_exprt inequality(
1581  effective_offset, ID_ge, std::move(zero));
1582 
1584  inequality,
1585  name + " lower bound",
1586  "array bounds",
1587  expr.find_source_location(),
1588  expr,
1589  guard);
1590  }
1591  }
1592  }
1593 
1594  if(ode.root_object().id() == ID_dereference)
1595  {
1596  const exprt &pointer = to_dereference_expr(ode.root_object()).pointer();
1597 
1598  const plus_exprt effective_offset{
1599  ode.offset(),
1601  pointer_offset(pointer), ode.offset().type())};
1602 
1603  binary_relation_exprt inequality{
1604  effective_offset,
1605  ID_lt,
1607  object_size(pointer), effective_offset.type())};
1608 
1609  exprt in_bounds_of_some_explicit_allocation =
1611  pointer,
1612  plus_exprt{ode.offset(), from_integer(1, ode.offset().type())});
1613 
1614  or_exprt precond(
1615  std::move(in_bounds_of_some_explicit_allocation), inequality);
1616 
1618  precond,
1619  name + " dynamic object upper bound",
1620  "array bounds",
1621  expr.find_source_location(),
1622  expr,
1623  guard);
1624 
1625  return;
1626  }
1627 
1628  const exprt &size = array_type.id() == ID_array
1629  ? to_array_type(array_type).size()
1630  : to_vector_type(array_type).size();
1631 
1632  if(size.is_nil() && !array_type.get_bool(ID_C_flexible_array_member))
1633  {
1634  // Linking didn't complete, we don't have a size.
1635  // Not clear what to do.
1636  }
1637  else if(size.id() == ID_infinity)
1638  {
1639  }
1640  else if(
1641  expr.array().id() == ID_member &&
1642  (size.is_zero() || array_type.get_bool(ID_C_flexible_array_member)))
1643  {
1644  // a variable sized struct member
1645  //
1646  // Excerpt from the C standard on flexible array members:
1647  // However, when a . (or ->) operator has a left operand that is (a pointer
1648  // to) a structure with a flexible array member and the right operand names
1649  // that member, it behaves as if that member were replaced with the longest
1650  // array (with the same element type) that would not make the structure
1651  // larger than the object being accessed; [...]
1652  const auto type_size_opt =
1654  CHECK_RETURN(type_size_opt.has_value());
1655 
1656  binary_relation_exprt inequality(
1657  ode.offset(),
1658  ID_lt,
1659  from_integer(type_size_opt.value(), ode.offset().type()));
1660 
1662  inequality,
1663  name + " upper bound",
1664  "array bounds",
1665  expr.find_source_location(),
1666  expr,
1667  guard);
1668  }
1669  else
1670  {
1671  binary_relation_exprt inequality{
1672  typecast_exprt::conditional_cast(index, size.type()), ID_lt, size};
1673 
1675  inequality,
1676  name + " upper bound",
1677  "array bounds",
1678  expr.find_source_location(),
1679  expr,
1680  guard);
1681  }
1682 }
1683 
1685  const unary_exprt &expr,
1686  const guardt &guard)
1687 {
1688  std::string name;
1689 
1690  if(expr.id() == ID_count_leading_zeros)
1691  name = "leading";
1692  else if(expr.id() == ID_count_trailing_zeros)
1693  name = "trailing";
1694  else
1695  PRECONDITION(false);
1696 
1698  notequal_exprt{expr.op(), from_integer(0, expr.op().type())},
1699  "count " + name + " zeros is undefined for value zero",
1700  "bit count",
1701  expr.find_source_location(),
1702  expr,
1703  guard);
1704 }
1705 
1707  const exprt &asserted_expr,
1708  const std::string &comment,
1709  const std::string &property_class,
1710  const source_locationt &source_location,
1711  const exprt &src_expr,
1712  const guardt &guard)
1713 {
1714  // first try simplifier on it
1715  exprt simplified_expr =
1716  enable_simplify ? simplify_expr(asserted_expr, ns) : asserted_expr;
1717 
1718  // throw away trivial properties?
1719  if(!retain_trivial && simplified_expr.is_true())
1720  return;
1721 
1722  // add the guard
1723  exprt guarded_expr = guard(simplified_expr);
1724 
1725  if(assertions.insert(std::make_pair(src_expr, guarded_expr)).second)
1726  {
1727  auto t = new_code.add(
1729  std::move(guarded_expr), source_location)
1731  std::move(guarded_expr), source_location));
1732 
1733  std::string source_expr_string;
1734  get_language_from_mode(mode)->from_expr(src_expr, source_expr_string, ns);
1735 
1736  t->source_location_nonconst().set_comment(
1737  comment + " in " + source_expr_string);
1738  t->source_location_nonconst().set_property_class(property_class);
1739 
1740  add_all_disable_named_check_pragmas(t->source_location_nonconst());
1741  }
1742 }
1743 
1744 void goto_check_ct::check_rec_address(const exprt &expr, const guardt &guard)
1745 {
1746  // we don't look into quantifiers
1747  if(expr.id() == ID_exists || expr.id() == ID_forall)
1748  return;
1749 
1750  if(expr.id() == ID_dereference)
1751  {
1752  check_rec(to_dereference_expr(expr).pointer(), guard);
1753  }
1754  else if(expr.id() == ID_index)
1755  {
1756  const index_exprt &index_expr = to_index_expr(expr);
1757  check_rec_address(index_expr.array(), guard);
1758  check_rec(index_expr.index(), guard);
1759  }
1760  else
1761  {
1762  for(const auto &operand : expr.operands())
1763  check_rec_address(operand, guard);
1764  }
1765 }
1766 
1767 void goto_check_ct::check_rec_logical_op(const exprt &expr, const guardt &guard)
1768 {
1769  PRECONDITION(
1770  expr.id() == ID_and || expr.id() == ID_or || expr.id() == ID_implies);
1771  INVARIANT(
1772  expr.is_boolean(),
1773  "'" + expr.id_string() + "' must be Boolean, but got " + expr.pretty());
1774 
1775  exprt::operandst constraints;
1776 
1777  for(const auto &op : expr.operands())
1778  {
1779  INVARIANT(
1780  op.is_boolean(),
1781  "'" + expr.id_string() + "' takes Boolean operands only, but got " +
1782  op.pretty());
1783 
1784  auto new_guard = [&guard, &constraints](exprt expr) {
1785  return guard(implication(conjunction(constraints), expr));
1786  };
1787 
1788  check_rec(op, new_guard);
1789 
1790  constraints.push_back(expr.id() == ID_or ? boolean_negate(op) : op);
1791  }
1792 }
1793 
1794 void goto_check_ct::check_rec_if(const if_exprt &if_expr, const guardt &guard)
1795 {
1796  INVARIANT(
1797  if_expr.cond().is_boolean(),
1798  "first argument of if must be boolean, but got " + if_expr.cond().pretty());
1799 
1800  check_rec(if_expr.cond(), guard);
1801 
1802  {
1803  auto new_guard = [&guard, &if_expr](exprt expr) {
1804  return guard(implication(if_expr.cond(), std::move(expr)));
1805  };
1806  check_rec(if_expr.true_case(), new_guard);
1807  }
1808 
1809  {
1810  auto new_guard = [&guard, &if_expr](exprt expr) {
1811  return guard(implication(not_exprt(if_expr.cond()), std::move(expr)));
1812  };
1813  check_rec(if_expr.false_case(), new_guard);
1814  }
1815 }
1816 
1818  const member_exprt &member,
1819  const guardt &guard)
1820 {
1821  const dereference_exprt &deref = to_dereference_expr(member.struct_op());
1822 
1823  check_rec(deref.pointer(), guard);
1824 
1825  // avoid building the following expressions when pointer_validity_check
1826  // would return immediately anyway
1828  return true;
1829 
1830  // we rewrite s->member into *(s+member_offset)
1831  // to avoid requiring memory safety of the entire struct
1832  auto member_offset_opt = member_offset_expr(member, ns);
1833 
1834  if(member_offset_opt.has_value())
1835  {
1836  pointer_typet new_pointer_type = to_pointer_type(deref.pointer().type());
1837  new_pointer_type.base_type() = member.type();
1838 
1839  const exprt char_pointer = typecast_exprt::conditional_cast(
1840  deref.pointer(), pointer_type(char_type()));
1841 
1842  const exprt new_address_casted = typecast_exprt::conditional_cast(
1843  plus_exprt{
1844  char_pointer,
1846  member_offset_opt.value(), pointer_diff_type())},
1847  new_pointer_type);
1848 
1849  dereference_exprt new_deref{new_address_casted};
1850  new_deref.add_source_location() = deref.source_location();
1851  pointer_validity_check(new_deref, member, guard);
1852 
1853  return true;
1854  }
1855  return false;
1856 }
1857 
1859  const div_exprt &div_expr,
1860  const guardt &guard)
1861 {
1862  div_by_zero_check(to_div_expr(div_expr), guard);
1863 
1864  if(div_expr.type().id() == ID_signedbv)
1865  integer_overflow_check(div_expr, guard);
1866  else if(div_expr.type().id() == ID_floatbv)
1867  {
1868  nan_check(div_expr, guard);
1869  float_overflow_check(div_expr, guard);
1870  }
1871 }
1872 
1874  const exprt &expr,
1875  const guardt &guard)
1876 {
1877  if(expr.type().id() == ID_signedbv || expr.type().id() == ID_unsignedbv)
1878  {
1879  integer_overflow_check(expr, guard);
1880 
1881  if(
1882  expr.operands().size() == 2 && expr.id() == ID_minus &&
1883  expr.operands()[0].type().id() == ID_pointer &&
1884  expr.operands()[1].type().id() == ID_pointer)
1885  {
1886  pointer_rel_check(to_binary_expr(expr), guard);
1887  }
1888  }
1889  else if(expr.type().id() == ID_floatbv)
1890  {
1891  nan_check(expr, guard);
1892  float_overflow_check(expr, guard);
1893  }
1894  else if(expr.type().id() == ID_pointer)
1895  {
1896  pointer_overflow_check(expr, guard);
1897  }
1898 }
1899 
1900 void goto_check_ct::check_rec(const exprt &expr, const guardt &guard)
1901 {
1902  if(expr.id() == ID_exists || expr.id() == ID_forall)
1903  {
1904  // the scoped variables may be used in the assertion
1905  const auto &quantifier_expr = to_quantifier_expr(expr);
1906 
1907  auto new_guard = [&guard, &quantifier_expr](exprt expr) {
1908  return guard(forall_exprt(quantifier_expr.symbol(), expr));
1909  };
1910 
1911  check_rec(quantifier_expr.where(), new_guard);
1912  return;
1913  }
1914  else if(expr.id() == ID_address_of)
1915  {
1916  check_rec_address(to_address_of_expr(expr).object(), guard);
1917  return;
1918  }
1919  else if(expr.id() == ID_and || expr.id() == ID_or || expr.id() == ID_implies)
1920  {
1921  check_rec_logical_op(expr, guard);
1922  return;
1923  }
1924  else if(expr.id() == ID_if)
1925  {
1926  check_rec_if(to_if_expr(expr), guard);
1927  return;
1928  }
1929  else if(
1930  expr.id() == ID_member &&
1931  to_member_expr(expr).struct_op().id() == ID_dereference)
1932  {
1933  if(check_rec_member(to_member_expr(expr), guard))
1934  return;
1935  }
1936 
1937  forall_operands(it, expr)
1938  check_rec(*it, guard);
1939 
1940  if(expr.type().id() == ID_c_enum_tag)
1941  enum_range_check(expr, guard);
1942 
1943  if(expr.id() == ID_index)
1944  {
1945  bounds_check(expr, guard);
1946  }
1947  else if(expr.id() == ID_div)
1948  {
1949  check_rec_div(to_div_expr(expr), guard);
1950  }
1951  else if(expr.id() == ID_shl || expr.id() == ID_ashr || expr.id() == ID_lshr)
1952  {
1953  undefined_shift_check(to_shift_expr(expr), guard);
1954 
1955  if(expr.id() == ID_shl && expr.type().id() == ID_signedbv)
1956  integer_overflow_check(expr, guard);
1957  }
1958  else if(expr.id() == ID_mod)
1959  {
1960  mod_by_zero_check(to_mod_expr(expr), guard);
1961  mod_overflow_check(to_mod_expr(expr), guard);
1962  }
1963  else if(
1964  expr.id() == ID_plus || expr.id() == ID_minus || expr.id() == ID_mult ||
1965  expr.id() == ID_unary_minus)
1966  {
1967  check_rec_arithmetic_op(expr, guard);
1968  }
1969  else if(expr.id() == ID_typecast)
1970  conversion_check(expr, guard);
1971  else if(
1972  expr.id() == ID_le || expr.id() == ID_lt || expr.id() == ID_ge ||
1973  expr.id() == ID_gt)
1975  else if(expr.id() == ID_dereference)
1976  {
1977  pointer_validity_check(to_dereference_expr(expr), expr, guard);
1978  }
1979  else if(requires_pointer_primitive_check(expr))
1980  {
1981  pointer_primitive_check(expr, guard);
1982  }
1983  else if(
1984  expr.id() == ID_count_leading_zeros || expr.id() == ID_count_trailing_zeros)
1985  {
1986  bounds_check(expr, guard);
1987  }
1988 }
1989 
1990 void goto_check_ct::check(const exprt &expr)
1991 {
1992  check_rec(expr, identity);
1993 }
1994 
1997 {
1998  bool modified = false;
1999 
2000  for(auto &op : expr.operands())
2001  {
2002  auto op_result = rw_ok_check(op);
2003  if(op_result.has_value())
2004  {
2005  op = *op_result;
2006  modified = true;
2007  }
2008  }
2009 
2010  if(expr.id() == ID_r_ok || expr.id() == ID_w_ok || expr.id() == ID_rw_ok)
2011  {
2012  // these get an address as first argument and a size as second
2014  expr.operands().size() == 2, "r/w_ok must have two operands");
2015 
2016  const auto conditions = get_pointer_dereferenceable_conditions(
2017  to_r_or_w_ok_expr(expr).pointer(), to_r_or_w_ok_expr(expr).size());
2018 
2019  exprt::operandst conjuncts;
2020 
2021  for(const auto &c : conditions)
2022  conjuncts.push_back(c.assertion);
2023 
2024  exprt c = conjunction(conjuncts);
2025  if(enable_simplify)
2026  simplify(c, ns);
2027  return c;
2028  }
2029  else if(modified)
2030  return std::move(expr);
2031  else
2032  return {};
2033 }
2034 
2036  const irep_idt &function_identifier,
2037  goto_functiont &goto_function)
2038 {
2039  const auto &function_symbol = ns.lookup(function_identifier);
2040  mode = function_symbol.mode;
2041 
2042  if(mode != ID_C && mode != ID_cpp)
2043  return;
2044 
2045  assertions.clear();
2046 
2047  bool did_something = false;
2048 
2050  util_make_unique<local_bitvector_analysist>(goto_function, ns);
2051 
2052  goto_programt &goto_program = goto_function.body;
2053 
2054  Forall_goto_program_instructions(it, goto_program)
2055  {
2056  current_target = it;
2057  goto_programt::instructiont &i = *it;
2058 
2059  flag_overridet resetter(i.source_location());
2060  const auto &pragmas = i.source_location().get_pragmas();
2061  for(const auto &d : pragmas)
2062  {
2063  // match named-check related pragmas
2064  auto matched = match_named_check(d.first);
2065  if(matched.has_value())
2066  {
2067  auto named_check = matched.value();
2068  auto name = named_check.first;
2069  auto status = named_check.second;
2070  bool *flag = name_to_flag.find(name)->second;
2071  switch(status)
2072  {
2073  case check_statust::ENABLE:
2074  resetter.set_flag(*flag, true, name);
2075  break;
2076  case check_statust::DISABLE:
2077  resetter.set_flag(*flag, false, name);
2078  break;
2079  case check_statust::CHECKED:
2080  resetter.disable_flag(*flag, name);
2081  break;
2082  }
2083  }
2084  }
2085 
2086  // add checked pragmas for all active checks
2088 
2089  new_code.clear();
2090 
2091  // we clear all recorded assertions if
2092  // 1) we want to generate all assertions or
2093  // 2) the instruction is a branch target
2094  if(retain_trivial || i.is_target())
2095  assertions.clear();
2096 
2097  if(i.has_condition())
2098  {
2099  check(i.condition());
2100  }
2101 
2102  // magic ERROR label?
2103  for(const auto &label : error_labels)
2104  {
2105  if(std::find(i.labels.begin(), i.labels.end(), label) != i.labels.end())
2106  {
2107  auto t = new_code.add(
2111  false_exprt{}, i.source_location()));
2112 
2113  t->source_location_nonconst().set_property_class("error label");
2114  t->source_location_nonconst().set_comment("error label " + label);
2115  t->source_location_nonconst().set("user-provided", true);
2116  }
2117  }
2118 
2119  if(i.is_other())
2120  {
2121  const auto &code = i.get_other();
2122  const irep_idt &statement = code.get_statement();
2123 
2124  if(statement == ID_expression)
2125  {
2126  check(code);
2127  }
2128  else if(statement == ID_printf)
2129  {
2130  for(const auto &op : code.operands())
2131  check(op);
2132  }
2133  }
2134  else if(i.is_assign())
2135  {
2136  const exprt &assign_lhs = i.assign_lhs();
2137  const exprt &assign_rhs = i.assign_rhs();
2138 
2139  // Disable enum range checks for left-hand sides as their values are yet
2140  // to be set (by this assignment).
2141  {
2142  flag_overridet resetter(i.source_location());
2143  resetter.disable_flag(enable_enum_range_check, "enum_range_check");
2144  check(assign_lhs);
2145  }
2146 
2147  check(assign_rhs);
2148 
2149  // the LHS might invalidate any assertion
2150  invalidate(assign_lhs);
2151  }
2152  else if(i.is_function_call())
2153  {
2154  // Disable enum range checks for left-hand sides as their values are yet
2155  // to be set (by this function call).
2156  {
2157  flag_overridet resetter(i.source_location());
2158  resetter.disable_flag(enable_enum_range_check, "enum_range_check");
2159  check(i.call_lhs());
2160  }
2161  check(i.call_function());
2162 
2163  for(const auto &arg : i.call_arguments())
2164  check(arg);
2165 
2166  // the call might invalidate any assertion
2167  assertions.clear();
2168  }
2169  else if(i.is_set_return_value())
2170  {
2171  check(i.return_value());
2172  // the return value invalidate any assertion
2173  invalidate(i.return_value());
2174  }
2175  else if(i.is_throw())
2176  {
2177  // this has no successor
2178  assertions.clear();
2179  }
2180  else if(i.is_dead())
2181  {
2183  {
2184  const symbol_exprt &variable = i.dead_symbol();
2185 
2186  // is it dirty?
2187  if(local_bitvector_analysis->dirty(variable))
2188  {
2189  // need to mark the dead variable as dead
2190  exprt lhs = ns.lookup(CPROVER_PREFIX "dead_object").symbol_expr();
2191  exprt address_of_expr = typecast_exprt::conditional_cast(
2192  address_of_exprt(variable), lhs.type());
2193  if_exprt rhs(
2195  std::move(address_of_expr),
2196  lhs);
2199  std::move(lhs), std::move(rhs), i.source_location()));
2200  t->code_nonconst().add_source_location() = i.source_location();
2201  }
2202  }
2203  }
2204  else if(i.is_end_function())
2205  {
2206  if(
2207  function_identifier == goto_functionst::entry_point() &&
2209  {
2210  const symbolt &leak = ns.lookup(CPROVER_PREFIX "memory_leak");
2211  const symbol_exprt leak_expr = leak.symbol_expr();
2212 
2213  // add self-assignment to get helpful counterexample output
2214  new_code.add(goto_programt::make_assignment(leak_expr, leak_expr));
2215 
2216  source_locationt source_location;
2217  source_location.set_function(function_identifier);
2218 
2219  equal_exprt eq(
2220  leak_expr, null_pointer_exprt(to_pointer_type(leak.type)));
2222  eq,
2223  "dynamically allocated memory never freed",
2224  "memory-leak",
2225  source_location,
2226  eq,
2227  identity);
2228  }
2229  }
2230 
2231  i.transform([this](exprt expr) { return rw_ok_check(expr); });
2232 
2233  for(auto &instruction : new_code.instructions)
2234  {
2235  if(instruction.source_location().is_nil())
2236  {
2237  instruction.source_location_nonconst().id(irep_idt());
2238 
2239  if(!it->source_location().get_file().empty())
2240  instruction.source_location_nonconst().set_file(
2241  it->source_location().get_file());
2242 
2243  if(!it->source_location().get_line().empty())
2244  instruction.source_location_nonconst().set_line(
2245  it->source_location().get_line());
2246 
2247  if(!it->source_location().get_function().empty())
2248  instruction.source_location_nonconst().set_function(
2249  it->source_location().get_function());
2250 
2251  if(!it->source_location().get_column().empty())
2252  {
2253  instruction.source_location_nonconst().set_column(
2254  it->source_location().get_column());
2255  }
2256  }
2257  }
2258 
2259  // insert new instructions -- make sure targets are not moved
2260  did_something |= !new_code.instructions.empty();
2261 
2262  while(!new_code.instructions.empty())
2263  {
2264  goto_program.insert_before_swap(it, new_code.instructions.front());
2265  new_code.instructions.pop_front();
2266  it++;
2267  }
2268  }
2269 
2270  if(did_something)
2271  remove_skip(goto_program);
2272 }
2273 
2276  const exprt &address,
2277  const exprt &size)
2278 {
2280  PRECONDITION(address.type().id() == ID_pointer);
2283 
2284  conditionst conditions;
2285 
2286  const exprt in_bounds_of_some_explicit_allocation =
2288 
2289  const bool unknown = flags.is_unknown() || flags.is_uninitialized();
2290 
2291  if(unknown)
2292  {
2293  conditions.push_back(conditiont{
2294  not_exprt{is_invalid_pointer_exprt{address}}, "pointer invalid"});
2295  }
2296 
2297  if(unknown || flags.is_dynamic_heap())
2298  {
2299  conditions.push_back(conditiont(
2300  or_exprt(
2301  in_bounds_of_some_explicit_allocation,
2302  not_exprt(deallocated(address, ns))),
2303  "deallocated dynamic object"));
2304  }
2305 
2306  if(unknown || flags.is_dynamic_local())
2307  {
2308  conditions.push_back(conditiont(
2309  or_exprt(
2310  in_bounds_of_some_explicit_allocation,
2311  not_exprt(dead_object(address, ns))),
2312  "dead object"));
2313  }
2314 
2315  if(flags.is_dynamic_heap())
2316  {
2317  const or_exprt object_bounds_violation(
2318  object_lower_bound(address, nil_exprt()),
2319  object_upper_bound(address, size));
2320 
2321  conditions.push_back(conditiont(
2322  or_exprt(
2323  in_bounds_of_some_explicit_allocation,
2324  not_exprt(object_bounds_violation)),
2325  "pointer outside dynamic object bounds"));
2326  }
2327 
2328  if(unknown || flags.is_dynamic_local() || flags.is_static_lifetime())
2329  {
2330  const or_exprt object_bounds_violation(
2331  object_lower_bound(address, nil_exprt()),
2332  object_upper_bound(address, size));
2333 
2334  conditions.push_back(conditiont(
2335  or_exprt(
2336  in_bounds_of_some_explicit_allocation,
2337  not_exprt(object_bounds_violation)),
2338  "pointer outside object bounds"));
2339  }
2340 
2341  if(unknown || flags.is_integer_address())
2342  {
2343  conditions.push_back(conditiont(
2344  implies_exprt(
2345  integer_address(address), in_bounds_of_some_explicit_allocation),
2346  "invalid integer address"));
2347  }
2348 
2349  return conditions;
2350 }
2351 
2354  const exprt &address,
2355  const exprt &size)
2356 {
2358  PRECONDITION(address.type().id() == ID_pointer);
2361 
2362  if(flags.is_unknown() || flags.is_uninitialized() || flags.is_null())
2363  {
2364  return {conditiont{
2365  or_exprt{
2367  not_exprt(null_pointer(address))},
2368  "pointer NULL"}};
2369  }
2370 
2371  return {};
2372 }
2373 
2375  const exprt &pointer,
2376  const exprt &size)
2377 {
2378  exprt::operandst alloc_disjuncts;
2379  for(const auto &a : allocations)
2380  {
2381  typecast_exprt int_ptr(pointer, a.first.type());
2382 
2383  binary_relation_exprt lb_check(a.first, ID_le, int_ptr);
2384 
2385  plus_exprt upper_bound{
2386  int_ptr, typecast_exprt::conditional_cast(size, int_ptr.type())};
2387 
2388  binary_relation_exprt ub_check{
2389  std::move(upper_bound), ID_le, plus_exprt{a.first, a.second}};
2390 
2391  alloc_disjuncts.push_back(
2392  and_exprt{std::move(lb_check), std::move(ub_check)});
2393  }
2394  return disjunction(alloc_disjuncts);
2395 }
2396 
2398  const irep_idt &function_identifier,
2399  goto_functionst::goto_functiont &goto_function,
2400  const namespacet &ns,
2401  const optionst &options,
2402  message_handlert &message_handler)
2403 {
2404  goto_check_ct goto_check(ns, options, message_handler);
2405  goto_check.goto_check(function_identifier, goto_function);
2406 }
2407 
2409  const namespacet &ns,
2410  const optionst &options,
2411  goto_functionst &goto_functions,
2412  message_handlert &message_handler)
2413 {
2414  goto_check_ct goto_check(ns, options, message_handler);
2415 
2416  goto_check.collect_allocations(goto_functions);
2417 
2418  for(auto &gf_entry : goto_functions.function_map)
2419  {
2420  goto_check.goto_check(gf_entry.first, gf_entry.second);
2421  }
2422 }
2423 
2425  const optionst &options,
2426  goto_modelt &goto_model,
2427  message_handlert &message_handler)
2428 {
2429  const namespacet ns(goto_model.symbol_table);
2430  goto_check_c(ns, options, goto_model.goto_functions, message_handler);
2431 }
2432 
2434  source_locationt &source_location) const
2435 {
2436  for(const auto &entry : name_to_flag)
2437  if(*(entry.second))
2438  source_location.add_pragma("checked:" + id2string(entry.first));
2439 }
2440 
2442  source_locationt &source_location) const
2443 {
2444  for(const auto &entry : name_to_flag)
2445  source_location.add_pragma("disable:" + id2string(entry.first));
2446 }
2447 
2450 {
2451  auto s = id2string(named_check);
2452  auto col = s.find(":");
2453 
2454  if(col == std::string::npos)
2455  return {}; // separator not found
2456 
2457  auto name = s.substr(col + 1);
2458 
2459  if(name_to_flag.find(name) == name_to_flag.end())
2460  return {}; // name unknown
2461 
2462  check_statust status;
2463  if(!s.compare(0, 6, "enable"))
2464  status = check_statust::ENABLE;
2465  else if(!s.compare(0, 7, "disable"))
2466  status = check_statust::DISABLE;
2467  else if(!s.compare(0, 7, "checked"))
2468  status = check_statust::CHECKED;
2469  else
2470  return {}; // prefix unknow
2471 
2472  // success
2473  return std::pair<irep_idt, check_statust>{name, status};
2474 }
messaget
Class that provides messages with a built-in verbosity 'level'.
Definition: message.h:154
Forall_goto_program_instructions
#define Forall_goto_program_instructions(it, program)
Definition: goto_program.h:1234
guard_exprt
Definition: guard_expr.h:23
UNREACHABLE
#define UNREACHABLE
This should be used to mark dead code.
Definition: invariant.h:503
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
goto_functiont::body
goto_programt body
Definition: goto_function.h:26
to_r_or_w_ok_expr
const r_or_w_ok_exprt & to_r_or_w_ok_expr(const exprt &expr)
Definition: pointer_expr.h:764
array_name
std::string array_name(const namespacet &ns, const exprt &expr)
Definition: array_name.cpp:19
to_unsignedbv_type
const unsignedbv_typet & to_unsignedbv_type(const typet &type)
Cast a typet to an unsignedbv_typet.
Definition: bitvector_types.h:191
goto_check_ct::log
messaget log
Definition: goto_check_c.cpp:98
ieee_floatt
Definition: ieee_float.h:116
typecast_exprt::conditional_cast
static exprt conditional_cast(const exprt &expr, const typet &type)
Definition: std_expr.h:2025
goto_check_ct::collect_allocations
void collect_allocations(const goto_functionst &goto_functions)
Fill the list of allocations allocationst with <address, size> for every allocation instruction.
Definition: goto_check_c.cpp:423
goto_programt::instructiont::is_throw
bool is_throw() const
Definition: goto_program.h:467
goto_check_ct::undefined_shift_check
void undefined_shift_check(const shift_exprt &, const guardt &)
Definition: goto_check_c.cpp:545
to_unary_expr
const unary_exprt & to_unary_expr(const exprt &expr)
Cast an exprt to a unary_exprt.
Definition: std_expr.h:361
multi_ary_exprt
A base class for multi-ary expressions Associativity is not specified.
Definition: std_expr.h:856
c_enum_typet
The type of C enums.
Definition: c_types.h:216
has_subexpr
bool has_subexpr(const exprt &expr, const std::function< bool(const exprt &)> &pred)
returns true if the expression has a subexpression that satisfies pred
Definition: expr_util.cpp:139
source_locationt::as_string
std::string as_string() const
Definition: source_location.h:25
conjunction
exprt conjunction(const exprt::operandst &op)
1) generates a conjunction for two or more operands 2) for one operand, returns the operand 3) return...
Definition: std_expr.cpp:62
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
goto_check_ct::get_pointer_is_null_condition
optionalt< goto_check_ct::conditiont > get_pointer_is_null_condition(const exprt &address, const exprt &size)
Definition: goto_check_c.cpp:2353
goto_check_ct::check_rec_if
void check_rec_if(const if_exprt &if_expr, const guardt &guard)
Check an if expression: check the if-condition alone, and then check the true/false-cases with the gu...
Definition: goto_check_c.cpp:1794
goto_check_ct::nan_check
void nan_check(const exprt &, const guardt &)
Definition: goto_check_c.cpp:1183
goto_check_ct::check_rec_div
void check_rec_div(const div_exprt &div_expr, const guardt &guard)
Check that a division is valid: check for division by zero, overflow and NaN (for floating point numb...
Definition: goto_check_c.cpp:1858
local_bitvector_analysis.h
goto_check_ct::enable_pointer_check
bool enable_pointer_check
Definition: goto_check_c.cpp:258
arith_tools.h
object_descriptor_exprt::build
void build(const exprt &expr, const namespacet &ns)
Given an expression expr, attempt to find the underlying object it represents by skipping over type c...
Definition: pointer_expr.cpp:107
source_locationt::set_function
void set_function(const irep_idt &function)
Definition: source_location.h:120
goto_check_ct::enable_div_by_zero_check
bool enable_div_by_zero_check
Definition: goto_check_c.cpp:260
null_pointer
exprt null_pointer(const exprt &pointer)
Definition: pointer_predicates.cpp:100
goto_programt::instructiont::get_other
const goto_instruction_codet & get_other() const
Get the statement for OTHER.
Definition: goto_program.h:321
flag_overridet::flag_overridet
flag_overridet(const source_locationt &source_location)
Definition: goto_check_c.cpp:342
goto_programt::instructiont::is_other
bool is_other() const
Definition: goto_program.h:471
forall_exprt
A forall expression.
Definition: mathematical_expr.h:337
to_dereference_expr
const dereference_exprt & to_dereference_expr(const exprt &expr)
Cast an exprt to a dereference_exprt.
Definition: pointer_expr.h:716
optionst
Definition: options.h:22
goto_programt::instructiont::source_location_nonconst
source_locationt & source_location_nonconst()
Definition: goto_program.h:345
goto_check_ct::integer_overflow_check
void integer_overflow_check(const exprt &, const guardt &)
Definition: goto_check_c.cpp:829
CHECK_RETURN
#define CHECK_RETURN(CONDITION)
Definition: invariant.h:495
goto_check_ct::array_name
std::string array_name(const exprt &)
Definition: goto_check_c.cpp:1509
typet
The type of an expression, extends irept.
Definition: type.h:28
source_locationt::get_pragmas
const irept::named_subt & get_pragmas() const
Definition: source_location.h:189
goto_check_ct::get_pointer_dereferenceable_conditions
conditionst get_pointer_dereferenceable_conditions(const exprt &address, const exprt &size)
Definition: goto_check_c.cpp:1496
goto_check_ct::enable_assert_to_assume
bool enable_assert_to_assume
Definition: goto_check_c.cpp:271
irept::pretty
std::string pretty(unsigned indent=0, unsigned max_indent=0) const
Definition: irep.cpp:495
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
goto_check_ct::match_named_check
named_check_statust match_named_check(const irep_idt &named_check) const
Matches a named-check string of the form.
Definition: goto_check_c.cpp:2449
goto_check_ct::bounds_check
void bounds_check(const exprt &, const guardt &)
Definition: goto_check_c.cpp:1514
symbolt::type
typet type
Type of symbol.
Definition: symbol.h:31
dereference_exprt
Operator to dereference a pointer.
Definition: pointer_expr.h:659
vector_typet::size
const constant_exprt & size() const
Definition: std_types.cpp:269
local_bitvector_analysist::flagst::is_dynamic_heap
bool is_dynamic_heap() const
Definition: local_bitvector_analysis.h:130
if_exprt
The trinary if-then-else operator.
Definition: std_expr.h:2322
goto_check_ct::conditiont::description
std::string description
Definition: goto_check_c.cpp:170
remove_skip
void remove_skip(goto_programt &goto_program, goto_programt::targett begin, goto_programt::targett end)
remove unnecessary skip statements
Definition: remove_skip.cpp:87
c_enum_typet::memberst
std::vector< c_enum_membert > memberst
Definition: c_types.h:253
pointer_predicates.h
goto_programt::instructiont::is_dead
bool is_dead() const
Definition: goto_program.h:473
goto_check_ct::mod_overflow_check
void mod_overflow_check(const mod_exprt &, const guardt &)
check a mod expression for the case INT_MIN % -1
Definition: goto_check_c.cpp:634
goto_check_ct::assertionst
std::set< std::pair< exprt, exprt > > assertionst
Definition: goto_check_c.cpp:248
goto_check_ct::enable_signed_overflow_check
bool enable_signed_overflow_check
Definition: goto_check_c.cpp:262
object_descriptor_exprt
Split an expression into a base object and a (byte) offset.
Definition: pointer_expr.h:166
goto_check_ct::assertions
assertionst assertions
Definition: goto_check_c.cpp:249
prefix.h
goto_check_ct::ns
const namespacet & ns
Definition: goto_check_c.cpp:95
get_language_from_mode
std::unique_ptr< languaget > get_language_from_mode(const irep_idt &mode)
Get the language corresponding to the given mode.
Definition: mode.cpp:51
invariant.h
and_exprt
Boolean AND.
Definition: std_expr.h:2070
source_locationt::add_pragma
void add_pragma(const irep_idt &pragma)
Definition: source_location.h:184
goto_check_c.h
goto_programt::add
targett add(instructiont &&instruction)
Adds a given instruction at the end.
Definition: goto_program.h:709
goto_model.h
plus_exprt
The plus expression Associativity is not specified.
Definition: std_expr.h:946
goto_programt::instructiont::transform
void transform(std::function< optionalt< exprt >(exprt)>)
Apply given transformer to all expressions; no return value means no change needed.
Definition: goto_program.cpp:982
goto_check_ct::name_to_flag
std::map< irep_idt, bool * > name_to_flag
Maps a named-check name to the corresponding boolean flag.
Definition: goto_check_c.cpp:275
exprt
Base class for all expressions.
Definition: expr.h:55
goto_modelt
Definition: goto_model.h:25
binary_exprt::lhs
exprt & lhs()
Definition: std_expr.h:613
unary_exprt
Generic base class for unary expressions.
Definition: std_expr.h:313
binary_exprt
A base class for binary expressions.
Definition: std_expr.h:582
mode.h
goto_check_ct::enable_simplify
bool enable_simplify
Definition: goto_check_c.cpp:268
namespace_baset::follow_tag
const union_typet & follow_tag(const union_tag_typet &) const
Follow type tag of union type.
Definition: namespace.cpp:63
goto_check_ct::invalidate
void invalidate(const exprt &lhs)
Remove all assertions containing the symbol in lhs as well as all assertions containing dereference.
Definition: goto_check_c.cpp:466
options.h
irep_idt
dstringt irep_idt
Definition: irep.h:37
configt::cppt::cpp_standardt::CPP17
@ CPP17
bool_typet
The Boolean type.
Definition: std_types.h:35
messaget::eom
static eomt eom
Definition: message.h:297
optionst::value_listt
std::list< std::string > value_listt
Definition: options.h:25
goto_check_ct::check_rec_logical_op
void check_rec_logical_op(const exprt &expr, const guardt &guard)
Check a logical operation: check each operand in separation while extending the guarding condition as...
Definition: goto_check_c.cpp:1767
exprt::is_true
bool is_true() const
Return whether the expression is a constant representing true.
Definition: expr.cpp:34
goto_check_ct::named_check_statust
optionalt< std::pair< irep_idt, check_statust > > named_check_statust
optional (named-check, status) pair
Definition: goto_check_c.cpp:319
configt::ansi_c
struct configt::ansi_ct ansi_c
to_bitvector_type
const bitvector_typet & to_bitvector_type(const typet &type)
Cast a typet to a bitvector_typet.
Definition: bitvector_types.h:32
flag_overridet::disabled_flags
std::set< bool * > disabled_flags
Definition: goto_check_c.cpp:405
make_binary
exprt make_binary(const exprt &expr)
splits an expression with >=3 operands into nested binary expressions
Definition: expr_util.cpp:36
goto_functionst::function_map
function_mapt function_map
Definition: goto_functions.h:29
div_exprt
Division.
Definition: std_expr.h:1096
symbol_exprt
Expression to hold a symbol (variable)
Definition: std_expr.h:112
shift_exprt::op
exprt & op()
Definition: bitvector_expr.h:239
goto_check_ct::check_rec_address
void check_rec_address(const exprt &expr, const guardt &guard)
Check an address-of expression: if it is a dereference then check the pointer if it is an index then ...
Definition: goto_check_c.cpp:1744
equal_exprt
Equality.
Definition: std_expr.h:1305
flag_overridet::flags_to_reset
std::map< bool *, bool > flags_to_reset
Definition: goto_check_c.cpp:404
goto_check_ct::pointer_primitive_check
void pointer_primitive_check(const exprt &expr, const guardt &guard)
Generates VCCs to check that pointers passed to pointer primitives are either null or valid.
Definition: goto_check_c.cpp:1432
to_minus_expr
const minus_exprt & to_minus_expr(const exprt &expr)
Cast an exprt to a minus_exprt.
Definition: std_expr.h:1031
goto_programt::instructiont::call_arguments
const exprt::operandst & call_arguments() const
Get the arguments of a FUNCTION_CALL.
Definition: goto_program.h:307
goto_check_ct::CHECKED
@ CHECKED
Definition: goto_check_c.cpp:315
goto_programt::make_assignment
static instructiont make_assignment(const code_assignt &_code, const source_locationt &l=source_locationt::nil())
Create an assignment instruction.
Definition: goto_program.h:1056
goto_check_c
void goto_check_c(const irep_idt &function_identifier, goto_functionst::goto_functiont &goto_function, const namespacet &ns, const optionst &options, message_handlert &message_handler)
Definition: goto_check_c.cpp:2397
if_exprt::false_case
exprt & false_case()
Definition: std_expr.h:2360
notequal_exprt
Disequality.
Definition: std_expr.h:1364
unsignedbv_typet
Fixed-width bit-vector with unsigned binary interpretation.
Definition: bitvector_types.h:158
object_descriptor_exprt::root_object
static const exprt & root_object(const exprt &expr)
Definition: pointer_expr.cpp:168
ieee_float_spect
Definition: ieee_float.h:22
to_binary_expr
const binary_exprt & to_binary_expr(const exprt &expr)
Cast an exprt to a binary_exprt.
Definition: std_expr.h:660
r_or_w_ok_exprt::pointer
const exprt & pointer() const
Definition: pointer_expr.h:753
goto_check_ct::conditionst
std::list< conditiont > conditionst
Definition: goto_check_c.cpp:173
goto_check_ct::enable_float_overflow_check
bool enable_float_overflow_check
Definition: goto_check_c.cpp:267
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
goto_check_ct::new_code
goto_programt new_code
Definition: goto_check_c.cpp:247
exprt::type
typet & type()
Return the type of the expression.
Definition: expr.h:84
namespacet::lookup
bool lookup(const irep_idt &name, const symbolt *&symbol) const override
See documentation for namespace_baset::lookup().
Definition: namespace.cpp:138
goto_check_ct::enable_unsigned_overflow_check
bool enable_unsigned_overflow_check
Definition: goto_check_c.cpp:263
goto_check_ct::current_target
goto_programt::const_targett current_target
Definition: goto_check_c.cpp:97
goto_check_ct::check_rec_arithmetic_op
void check_rec_arithmetic_op(const exprt &expr, const guardt &guard)
Check that an arithmetic operation is valid: overflow check, NaN-check (for floating point numbers),...
Definition: goto_check_c.cpp:1873
disjunction
exprt disjunction(const exprt::operandst &op)
1) generates a disjunction for two or more operands 2) for one operand, returns the operand 3) return...
Definition: std_expr.cpp:50
goto_check_ct::goto_functiont
goto_functionst::goto_functiont goto_functiont
Definition: goto_check_c.cpp:81
goto_check_ct::rw_ok_check
optionalt< exprt > rw_ok_check(exprt)
expand the r_ok, w_ok and rw_ok predicates
Definition: goto_check_c.cpp:1996
goto_programt::make_assertion
static instructiont make_assertion(const exprt &g, const source_locationt &l=source_locationt::nil())
Definition: goto_program.h:924
make_unique.h
find_symbols.h
goto_check_ct::check
void check(const exprt &expr)
Initiate the recursively analysis of expr with its ‘guard’ set to TRUE.
Definition: goto_check_c.cpp:1990
has_prefix
bool has_prefix(const std::string &s, const std::string &prefix)
Definition: converter.cpp:13
or_exprt
Boolean OR.
Definition: std_expr.h:2178
to_mod_expr
const mod_exprt & to_mod_expr(const exprt &expr)
Cast an exprt to a mod_exprt.
Definition: std_expr.h:1217
goto_check_ct::enable_pointer_primitive_check
bool enable_pointer_primitive_check
Definition: goto_check_c.cpp:272
goto_check_ct::enum_range_check
void enum_range_check(const exprt &, const guardt &)
Definition: goto_check_c.cpp:517
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
can_cast_expr< object_size_exprt >
bool can_cast_expr< object_size_exprt >(const exprt &base)
Definition: pointer_expr.h:1049
null_pointer_exprt
The null pointer constant.
Definition: pointer_expr.h:734
goto_programt::instructiont::dead_symbol
const symbol_exprt & dead_symbol() const
Get the symbol for DEAD.
Definition: goto_program.h:251
id2string
const std::string & id2string(const irep_idt &d)
Definition: irep.h:47
goto_check_ct::float_overflow_check
void float_overflow_check(const exprt &, const guardt &)
Definition: goto_check_c.cpp:1075
goto_programt::instructiont::source_location
const source_locationt & source_location() const
Definition: goto_program.h:340
goto_check_ct::identity
const guardt identity
Definition: goto_check_c.cpp:101
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
goto_check_ct::conditiont::conditiont
conditiont(const exprt &_assertion, const std::string &_description)
Definition: goto_check_c.cpp:164
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
implication
static exprt implication(exprt lhs, exprt rhs)
Definition: goto_check_c.cpp:408
goto_check_ct::enable_undefined_shift_check
bool enable_undefined_shift_check
Definition: goto_check_c.cpp:266
binary_overflow_exprt
A Boolean expression returning true, iff operation kind would result in an overflow when applied to o...
Definition: bitvector_expr.h:704
member_exprt::struct_op
const exprt & struct_op() const
Definition: std_expr.h:2832
binary_predicate_exprt
A base class for expressions that are predicates, i.e., Boolean-typed, and that take exactly two argu...
Definition: std_expr.h:675
PRECONDITION
#define PRECONDITION(CONDITION)
Definition: invariant.h:463
exprt::find_source_location
const source_locationt & find_source_location() const
Get a source_locationt from the expression or from its operands (non-recursively).
Definition: expr.cpp:165
symbol_exprt::get_identifier
const irep_idt & get_identifier() const
Definition: std_expr.h:142
nil_exprt
The NIL expression.
Definition: std_expr.h:3025
dereference_exprt::pointer
exprt & pointer()
Definition: pointer_expr.h:673
boolean_negate
exprt boolean_negate(const exprt &src)
negate a Boolean expression, possibly removing a not_exprt, and swapping false and true
Definition: expr_util.cpp:127
flag_overridet::~flag_overridet
~flag_overridet()
Restore the values of all flags that have been modified via set_flag.
Definition: goto_check_c.cpp:396
object_upper_bound
exprt object_upper_bound(const exprt &pointer, const exprt &access_size)
Definition: pointer_predicates.cpp:106
goto_check_ct::local_bitvector_analysis
std::unique_ptr< local_bitvector_analysist > local_bitvector_analysis
Definition: goto_check_c.cpp:96
symbolt::symbol_expr
class symbol_exprt symbol_expr() const
Produces a symbol_exprt for a symbol.
Definition: symbol.cpp:121
local_bitvector_analysist::flagst::is_integer_address
bool is_integer_address() const
Definition: local_bitvector_analysis.h:160
to_plus_expr
const plus_exprt & to_plus_expr(const exprt &expr)
Cast an exprt to a plus_exprt.
Definition: std_expr.h:986
flag_overridet::source_location
const source_locationt & source_location
Definition: goto_check_c.cpp:403
pointer_expr.h
goto_check_ct::allocationst
std::list< allocationt > allocationst
Definition: goto_check_c.cpp:296
goto_programt::instructiont::labels
labelst labels
Definition: goto_program.h:419
goto_check_ct::conversion_check
void conversion_check(const exprt &, const guardt &)
Definition: goto_check_c.cpp:666
goto_check_ct::guardt
std::function< exprt(exprt)> guardt
Definition: goto_check_c.cpp:100
deallocated
exprt deallocated(const exprt &pointer, const namespacet &ns)
Definition: pointer_predicates.cpp:43
irept::id_string
const std::string & id_string() const
Definition: irep.h:399
simplify
bool simplify(exprt &expr, const namespacet &ns)
Definition: simplify_expr.cpp:2926
dead_object
exprt dead_object(const exprt &pointer, const namespacet &ns)
Definition: pointer_predicates.cpp:51
mult_exprt
Binary multiplication Associativity is not specified.
Definition: std_expr.h:1051
goto_check_ct::allocations
allocationst allocations
Definition: goto_check_c.cpp:297
language.h
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
integer_address
exprt integer_address(const exprt &pointer)
Definition: pointer_predicates.cpp:93
index_exprt::index
exprt & index()
Definition: std_expr.h:1450
goto_programt::instructiont::has_condition
bool has_condition() const
Does this instruction have a condition?
Definition: goto_program.h:367
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
goto_check_ct::pointer_validity_check
void pointer_validity_check(const dereference_exprt &expr, const exprt &src_expr, const guardt &guard)
Generates VCCs for the validity of the given dereferencing operation.
Definition: goto_check_c.cpp:1391
pointer_type
pointer_typet pointer_type(const typet &subtype)
Definition: c_types.cpp:253
goto_programt::instructiont::is_set_return_value
bool is_set_return_value() const
Definition: goto_program.h:464
index_exprt::array
exprt & array()
Definition: std_expr.h:1440
local_bitvector_analysist::flagst::is_null
bool is_null() const
Definition: local_bitvector_analysis.h:140
to_symbol_expr
const symbol_exprt & to_symbol_expr(const exprt &expr)
Cast an exprt to a symbol_exprt.
Definition: std_expr.h:222
goto_check_ct::check_rec
void check_rec(const exprt &expr, const guardt &guard)
Recursively descend into expr and run the appropriate check for each sub-expression,...
Definition: goto_check_c.cpp:1900
object_size
exprt object_size(const exprt &pointer)
Definition: pointer_predicates.cpp:33
local_bitvector_analysist::flagst::is_static_lifetime
bool is_static_lifetime() const
Definition: local_bitvector_analysis.h:150
irept::is_nil
bool is_nil() const
Definition: irep.h:376
irept::id
const irep_idt & id() const
Definition: irep.h:396
message_handlert
Definition: message.h:27
goto_programt::instructiont::is_end_function
bool is_end_function() const
Definition: goto_program.h:480
goto_functiont
A goto function, consisting of function body (see body) and parameter identifiers (see parameter_iden...
Definition: goto_function.h:23
exprt::operandst
std::vector< exprt > operandst
Definition: expr.h:58
code_function_callt::argumentst
exprt::operandst argumentst
Definition: goto_instruction_code.h:281
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
unary_exprt::op
const exprt & op() const
Definition: std_expr.h:326
to_shl_expr
const shl_exprt & to_shl_expr(const exprt &expr)
Cast an exprt to a shl_exprt.
Definition: bitvector_expr.h:320
goto_check_ct::enable_enum_range_check
bool enable_enum_range_check
Definition: goto_check_c.cpp:261
goto_programt::instructiont::call_lhs
const exprt & call_lhs() const
Get the lhs of a FUNCTION_CALL (may be nil)
Definition: goto_program.h:293
goto_programt::instructiont::assign_rhs
const exprt & assign_rhs() const
Get the rhs of the assignment for ASSIGN.
Definition: goto_program.h:221
goto_check_ct::enable_memory_leak_check
bool enable_memory_leak_check
Definition: goto_check_c.cpp:259
goto_check_ct::retain_trivial
bool retain_trivial
Definition: goto_check_c.cpp:270
cprover_prefix.h
local_bitvector_analysist::flagst::is_dynamic_local
bool is_dynamic_local() const
Definition: local_bitvector_analysis.h:120
ieee_floatt::from_integer
void from_integer(const mp_integer &i)
Definition: ieee_float.cpp:515
object_lower_bound
exprt object_lower_bound(const exprt &pointer, const exprt &offset)
Definition: pointer_predicates.cpp:136
std_code.h
optionalt
nonstd::optional< T > optionalt
Definition: optional.h:35
goto_check_ct::mod_by_zero_check
void mod_by_zero_check(const mod_exprt &, const guardt &)
Definition: goto_check_c.cpp:612
goto_check_ct::allocationt
std::pair< exprt, exprt > allocationt
Definition: goto_check_c.cpp:295
goto_check_ct::bounds_check_index
void bounds_check_index(const index_exprt &, const guardt &)
Definition: goto_check_c.cpp:1531
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
null_object
exprt null_object(const exprt &pointer)
Definition: pointer_predicates.cpp:87
goto_check_ct::pointer_overflow_check
void pointer_overflow_check(const exprt &, const guardt &)
Definition: goto_check_c.cpp:1335
goto_check_ct::get_pointer_points_to_valid_memory_conditions
conditionst get_pointer_points_to_valid_memory_conditions(const exprt &address, const exprt &size)
Definition: goto_check_c.cpp:2275
goto_programt::clear
void clear()
Clear the goto program.
Definition: goto_program.h:811
ieee_floatt::plus_infinity
static ieee_floatt plus_infinity(const ieee_float_spect &_spec)
Definition: ieee_float.h:212
goto_programt::instructiont::return_value
const exprt & return_value() const
Get the return value of a SET_RETURN_VALUE instruction.
Definition: goto_program.h:265
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
char_type
bitvector_typet char_type()
Definition: c_types.cpp:124
goto_check_ct::add_all_disable_named_check_pragmas
void add_all_disable_named_check_pragmas(source_locationt &source_location) const
Adds "disable" pragmas for all named checks on the given source location.
Definition: goto_check_c.cpp:2441
flag_overridet::set_flag
void set_flag(bool &flag, bool new_value, const irep_idt &flag_name)
Store the current value of flag and then set its value to new_value.
Definition: goto_check_c.cpp:352
source_locationt
Definition: source_location.h:18
configt::ansi_ct::c_standardt::C99
@ C99
simplify_expr.h
goto_check_ct::goto_check
void goto_check(const irep_idt &function_identifier, goto_functiont &goto_function)
Definition: goto_check_c.cpp:2035
bitvector_typet::get_width
std::size_t get_width() const
Definition: std_types.h:876
side_effect_expr_nondett
A side_effect_exprt that returns a non-deterministically chosen value.
Definition: std_code.h:1519
goto_functionst::goto_functiont
::goto_functiont goto_functiont
Definition: goto_functions.h:27
exprt::is_zero
bool is_zero() const
Return whether the expression is a constant representing 0.
Definition: expr.cpp:65
goto_check_ct::ENABLE
@ ENABLE
Definition: goto_check_c.cpp:313
local_bitvector_analysist::flagst::is_unknown
bool is_unknown() const
Definition: local_bitvector_analysis.h:90
member_exprt
Extract member of struct or union.
Definition: std_expr.h:2793
goto_programt::instructions
instructionst instructions
The list of instructions in the goto program.
Definition: goto_program.h:592
mod_exprt::divisor
exprt & divisor()
The divisor of a division is the number the dividend is being divided by.
Definition: std_expr.h:1188
expr_util.h
Deprecated expression utility functions.
configt::ansi_ct::c_standardt::C11
@ C11
goto_functionst
A collection of goto functions.
Definition: goto_functions.h:24
shift_exprt
A base class for shift and rotate operators.
Definition: bitvector_expr.h:229
goto_check_ct::enable_bounds_check
bool enable_bounds_check
Definition: goto_check_c.cpp:257
c_enum_tag_typet
C enum tag type, i.e., c_enum_typet with an identifier.
Definition: c_types.h:318
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
if_exprt::true_case
exprt & true_case()
Definition: std_expr.h:2350
goto_modelt::goto_functions
goto_functionst goto_functions
GOTO functions.
Definition: goto_model.h:33
to_quantifier_expr
const quantifier_exprt & to_quantifier_expr(const exprt &expr)
Cast an exprt to a quantifier_exprt.
Definition: mathematical_expr.h:319
unary_minus_overflow_exprt
A Boolean expression returning true, iff negation would result in an overflow when applied to the (si...
Definition: bitvector_expr.h:909
goto_programt::instructiont::assign_lhs
const exprt & assign_lhs() const
Get the lhs of the assignment for ASSIGN.
Definition: goto_program.h:207
shift_exprt::distance
exprt & distance()
Definition: bitvector_expr.h:249
goto_check_ct
Definition: goto_check_c.cpp:47
local_bitvector_analysist::flagst
Definition: local_bitvector_analysis.h:48
isinf_exprt
Evaluates to true if the operand is infinite.
Definition: floatbv_expr.h:131
goto_check_ct::add_guarded_property
void add_guarded_property(const exprt &asserted_expr, const std::string &comment, const std::string &property_class, const source_locationt &source_location, const exprt &src_expr, const guardt &guard)
Include the asserted_expr in the code conditioned by the guard.
Definition: goto_check_c.cpp:1706
to_implies_expr
const implies_exprt & to_implies_expr(const exprt &expr)
Cast an exprt to a implies_exprt.
Definition: std_expr.h:2159
symbolt
Symbol table entry.
Definition: symbol.h:27
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
configt::ansi_ct::c_standard
enum configt::ansi_ct::c_standardt c_standard
ieee_float.h
goto_check_ct::goto_check_ct
goto_check_ct(const namespacet &_ns, const optionst &_options, message_handlert &_message_handler)
Definition: goto_check_c.cpp:50
optionst::get_bool_option
bool get_bool_option(const std::string &option) const
Definition: options.cpp:44
goto_check_ct::bounds_check_bit_count
void bounds_check_bit_count(const unary_exprt &, const guardt &)
Definition: goto_check_c.cpp:1684
goto_check_ct::error_labelst
optionst::value_listt error_labelst
Definition: goto_check_c.cpp:290
flag_overridet
Allows to:
Definition: goto_check_c.cpp:339
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
has_symbol_expr
bool has_symbol_expr(const exprt &src, const irep_idt &identifier, bool include_bound_symbols)
Returns true if one of the symbol expressions in src has identifier identifier; if include_bound_symb...
Definition: find_symbols.cpp:232
goto_check_ct::is_in_bounds_of_some_explicit_allocation
exprt is_in_bounds_of_some_explicit_allocation(const exprt &pointer, const exprt &size)
Definition: goto_check_c.cpp:2374
goto_check_ct::conditiont
Definition: goto_check_c.cpp:162
goto_programt::instructiont::is_assign
bool is_assign() const
Definition: goto_program.h:465
pointer_typet::base_type
const typet & base_type() const
The type of the data what we point to.
Definition: pointer_expr.h:35
CPROVER_PREFIX
#define CPROVER_PREFIX
Definition: cprover_prefix.h:14
local_bitvector_analysist::flagst::is_uninitialized
bool is_uninitialized() const
Definition: local_bitvector_analysis.h:100
to_array_type
const array_typet & to_array_type(const typet &type)
Cast a typet to an array_typet.
Definition: std_types.h:844
same_object
exprt same_object(const exprt &p1, const exprt &p2)
Definition: pointer_predicates.cpp:28
flag_overridet::disable_flag
void disable_flag(bool &flag, const irep_idt &flag_name)
Sets the given flag to false, overriding any previous value.
Definition: goto_check_c.cpp:374
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
goto_programt
A generic container class for the GOTO intermediate representation of one function.
Definition: goto_program.h:72
configt::cppt::cpp_standardt::CPP14
@ CPP14
goto_check_ct::div_by_zero_check
void div_by_zero_check(const div_exprt &, const guardt &)
Definition: goto_check_c.cpp:496
config.h
goto_check_ct::check_rec_member
bool check_rec_member(const member_exprt &member, const guardt &guard)
Check that a member expression is valid:
Definition: goto_check_c.cpp:1817
to_vector_type
const vector_typet & to_vector_type(const typet &type)
Cast a typet to a vector_typet.
Definition: std_types.h:1060
size_of_expr
optionalt< exprt > size_of_expr(const typet &type, const namespacet &ns)
Definition: pointer_offset_size.cpp:280
to_member_expr
const member_exprt & to_member_expr(const exprt &expr)
Cast an exprt to a member_exprt.
Definition: std_expr.h:2886
goto_programt::const_targett
instructionst::const_iterator const_targett
Definition: goto_program.h:587
implies_exprt
Boolean implication.
Definition: std_expr.h:2133
ieee_floatt::to_expr
constant_exprt to_expr() const
Definition: ieee_float.cpp:702
goto_programt::instructiont::condition
const exprt & condition() const
Get the condition of gotos, assume, assert.
Definition: goto_program.h:373
configt::cppt::cpp_standardt::CPP11
@ CPP11
goto_programt::make_assumption
static instructiont make_assumption(const exprt &g, const source_locationt &l=source_locationt::nil())
Definition: goto_program.h:936
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
extractbits_exprt
Extracts a sub-range of a bit-vector operand.
Definition: bitvector_expr.h:447
exprt::operands
operandst & operands()
Definition: expr.h:94
goto_check_ct::check_statust
check_statust
activation statuses for named checks
Definition: goto_check_c.cpp:311
goto_functionst::entry_point
static irep_idt entry_point()
Get the identifier of the entry point to a goto model.
Definition: goto_functions.h:92
index_exprt
Array index operator.
Definition: std_expr.h:1409
is_invalid_pointer_exprt
Definition: pointer_predicates.h:37
address_of_exprt
Operator to return the address of an object.
Definition: pointer_expr.h:370
exprt::add_source_location
source_locationt & add_source_location()
Definition: expr.h:216
goto_programt::insert_before_swap
void insert_before_swap(targett target)
Insertion that preserves jumps to "target".
Definition: goto_program.h:613
configt::cpp
struct configt::cppt cpp
typecast_exprt
Semantic type conversion.
Definition: std_expr.h:2016
pointer_diff_type
signedbv_typet pointer_diff_type()
Definition: c_types.cpp:238
goto_check_ct::DISABLE
@ DISABLE
Definition: goto_check_c.cpp:314
pointer_typet
The pointer type These are both 'bitvector_typet' (they have a width) and 'type_with_subtypet' (they ...
Definition: pointer_expr.h:23
size_type
unsignedbv_typet size_type()
Definition: c_types.cpp:68
remove_skip.h
message.h
constant_exprt
A constant literal expression.
Definition: std_expr.h:2941
goto_check_ct::enable_conversion_check
bool enable_conversion_check
Definition: goto_check_c.cpp:265
ieee_float_equal_exprt
IEEE-floating-point equality.
Definition: floatbv_expr.h:263
goto_programt::instructiont::is_function_call
bool is_function_call() const
Definition: goto_program.h:466
exprt::is_boolean
bool is_boolean() const
Return whether the expression represents a Boolean.
Definition: expr.cpp:52
source_locationt::is_built_in
static bool is_built_in(const std::string &s)
Definition: source_location.cpp:16
messaget::warning
mstreamt & warning() const
Definition: message.h:404
multi_ary_exprt::op0
exprt & op0()
Definition: std_expr.h:877
member_offset_expr
optionalt< exprt > member_offset_expr(const member_exprt &member_expr, const namespacet &ns)
Definition: pointer_offset_size.cpp:221
array_name.h
binary_exprt::op1
exprt & op1()
Definition: expr.h:128
comment
static std::string comment(const rw_set_baset::entryt &entry, bool write)
Definition: race_check.cpp:109
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
goto_programt::instructiont
This class represents an instruction in the GOTO intermediate representation.
Definition: goto_program.h:180
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
goto_programt::instructiont::is_target
bool is_target() const
Is this node a branch target?
Definition: goto_program.h:425
constant_exprt::get_value
const irep_idt & get_value() const
Definition: std_expr.h:2950
exprt::source_location
const source_locationt & source_location() const
Definition: expr.h:211
goto_modelt::symbol_table
symbol_tablet symbol_table
Symbol table.
Definition: goto_model.h:30
goto_check_ct::add_active_named_check_pragmas
void add_active_named_check_pragmas(source_locationt &source_location) const
Adds "checked" pragmas for all currently active named checks on the given source location.
Definition: goto_check_c.cpp:2433
c_types.h
goto_check_ct::pointer_rel_check
void pointer_rel_check(const binary_exprt &, const guardt &)
Definition: goto_check_c.cpp:1291
goto_check_ct::enable_pointer_overflow_check
bool enable_pointer_overflow_check
Definition: goto_check_c.cpp:264
goto_check_ct::requires_pointer_primitive_check
bool requires_pointer_primitive_check(const exprt &expr)
Returns true if the given expression is a pointer primitive that requires validation of the operand t...
Definition: goto_check_c.cpp:1478
goto_programt::instructiont::call_function
const exprt & call_function() const
Get the function that is called for FUNCTION_CALL.
Definition: goto_program.h:279
optionst::get_list_option
const value_listt & get_list_option(const std::string &option) const
Definition: options.cpp:80
irept::get_bool
bool get_bool(const irep_idt &name) const
Definition: irep.cpp:58
ieee_floatt::minus_infinity
static ieee_floatt minus_infinity(const ieee_float_spect &_spec)
Definition: ieee_float.h:215
goto_check_ct::mode
irep_idt mode
Definition: goto_check_c.cpp:299
goto_programt::targett
instructionst::iterator targett
Definition: goto_program.h:586
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
binary_exprt::op0
exprt & op0()
Definition: expr.h:125
validation_modet::INVARIANT
@ INVARIANT
mod_exprt
Modulo defined as lhs-(rhs * truncate(lhs/rhs)).
Definition: std_expr.h:1167
c_enum_typet::members
const memberst & members() const
Definition: c_types.h:255
goto_check_ct::conditiont::assertion
exprt assertion
Definition: goto_check_c.cpp:169
goto_check_ct::error_labels
error_labelst error_labels
Definition: goto_check_c.cpp:291
shl_exprt
Left shift.
Definition: bitvector_expr.h:294
goto_check_ct::enable_nan_check
bool enable_nan_check
Definition: goto_check_c.cpp:269
configt::cppt::cpp_standard
enum configt::cppt::cpp_standardt cpp_standard
mathematical_expr.h
not_exprt
Boolean negation.
Definition: std_expr.h:2277
object_descriptor_exprt::offset
exprt & offset()
Definition: pointer_expr.h:209