CBMC
cegis_verifier.cpp
Go to the documentation of this file.
1 /*******************************************************************\
2 
3 Module: Verifier for Counterexample-Guided Synthesis
4 
5 Author: Qinheping Hu
6 
7 \*******************************************************************/
8 
11 
12 #include "cegis_verifier.h"
13 
14 #include <util/arith_tools.h>
15 #include <util/c_types.h>
16 #include <util/options.h>
19 
24 
26 #include <ansi-c/cprover_library.h>
29 #include <assembler/remove_asm.h>
30 #include <cpp/cprover_library.h>
37 #include <langapi/language_util.h>
39 #include <solvers/prop/prop.h>
40 
41 static bool contains_symbol_prefix(const exprt &expr, const std::string &prefix)
42 {
43  for(auto it = expr.depth_begin(), itend = expr.depth_end(); it != itend; ++it)
44  {
45  if(
46  it->id() == ID_symbol &&
48  {
49  return true;
50  }
51  }
52  return false;
53 }
54 
55 static const exprt &
57 {
58  // A NULL-pointer check is the negation of an equation between the checked
59  // pointer and a NULL pointer.
60  // ! (POINTER_OBJECT(NULL) == POINTER_OBJECT(ptr))
61  const equal_exprt &equal_expr = to_equal_expr(to_not_expr(violation).op());
62 
63  const pointer_object_exprt &lhs_pointer_object =
64  to_pointer_object_expr(equal_expr.lhs());
65  const pointer_object_exprt &rhs_pointer_object =
66  to_pointer_object_expr(equal_expr.rhs());
67 
68  const exprt &lhs_pointer = lhs_pointer_object.operands()[0];
69  const exprt &rhs_pointer = rhs_pointer_object.operands()[0];
70 
71  // NULL == ptr
72  if(
73  can_cast_expr<constant_exprt>(lhs_pointer) &&
74  is_null_pointer(*expr_try_dynamic_cast<constant_exprt>(lhs_pointer)))
75  {
76  return rhs_pointer;
77  }
78 
79  // Not a equation with NULL on one side.
81 }
82 
84 {
85  // Preprocess `goto_model`. Copied from `cbmc_parse_options.cpp`.
91  // library functions may introduce inline assembler
92  while(has_asm(goto_model))
93  {
99  }
101 
103 
106 }
107 
109 cegis_verifiert::extract_violation_type(const std::string &description)
110 {
111  // The violation is a pointer OOB check.
112  if((description.find(
113  "dereference failure: pointer outside object bounds in") !=
114  std::string::npos))
115  {
117  }
118 
119  // The violation is a null pointer check.
120  if(description.find("pointer NULL") != std::string::npos)
121  {
123  }
124 
125  // The violation is a loop-invariant-preservation check.
126  if(description.find("preserved") != std::string::npos)
127  {
129  }
130 
131  // The violation is a loop-invariant-preservation check.
132  if(description.find("invariant before entry") != std::string::npos)
133  {
135  }
136 
137  // The violation is an assignable check.
138  if(description.find("assignable") != std::string::npos)
139  {
141  }
142 
144 }
145 
146 std::list<loop_idt>
148 {
149  std::list<loop_idt> result;
150 
151  // We say a loop is the cause loop of an assignable-violation if the inclusion
152  // check is in the loop.
153 
154  // So we check what loops the last step of the trace is in.
155  // Transformed loop head:
156  // ASSIGN entered_loop = false
157  // Transformed loop end:
158  // ASSIGN entered_loop = true
159  for(const auto &step : goto_trace.steps)
160  {
161  // We are entering a loop.
162  if(is_transformed_loop_head(step.pc))
163  {
164  result.push_front(
165  loop_idt(step.function_id, original_loop_number_map[step.pc]));
166  }
167  // We are leaving a loop.
168  else if(is_transformed_loop_end(step.pc))
169  {
170  const loop_idt loop_id(
171  step.function_id, original_loop_number_map[step.pc]);
172  INVARIANT(
173  result.front() == loop_id, "Leaving a loop we haven't entered.");
174  result.pop_front();
175  }
176  }
177 
178  INVARIANT(!result.empty(), "The assignable violation is not in a loop.");
179  return result;
180 }
181 
183  const goto_tracet &goto_trace,
184  const goto_programt::const_targett violation)
185 {
186  std::list<loop_idt> result;
187 
188  // build the dependence graph
189  dependence_grapht dependence_graph(ns, log.get_message_handler());
190  dependence_graph(goto_model);
191 
192  // Checking if `to` is dependent on `from`.
193  // `from` : loop havocing
194  // `to` : violation
195 
196  // Get `to`---the instruction where the violation happens
197  irep_idt to_fun_name = goto_trace.get_last_step().function_id;
198  const goto_functionst::goto_functiont &to_function =
199  goto_model.get_goto_function(to_fun_name);
200  goto_programt::const_targett to = to_function.body.instructions.end();
201  for(goto_programt::const_targett it = to_function.body.instructions.begin();
202  it != to_function.body.instructions.end();
203  ++it)
204  {
205  if(it->location_number == violation->location_number)
206  {
207  to = it;
208  }
209  }
210 
211  INVARIANT(
212  to != to_function.body.instructions.end(),
213  "There must be a violation in a trace.");
214 
215  // Compute the backward reachable set.
216  const auto reachable_vector =
217  dependence_graph.get_reachable(dependence_graph[to].get_node_id(), false);
218  const std::set<size_t> reachable_set =
219  std::set<size_t>(reachable_vector.begin(), reachable_vector.end());
220 
221  // A loop is the cause loop if the violation is dependent on the write set of
222  // the loop.
223  for(const auto &step : goto_trace.steps)
224  {
225  // Being dependent on a write set is equivalent to being dependent on one
226  // of the loop havocing instruction.
227  if(loop_havoc_set.count(step.pc))
228  {
229  // Get `from`---a loop havoc instruction.
230  irep_idt from_fun_name = step.function_id;
231  const goto_functionst::goto_functiont &from_function =
232  goto_model.get_goto_function(from_fun_name);
233  goto_programt::const_targett from = from_function.body.instructions.end();
235  from_function.body.instructions.begin();
236  it != from_function.body.instructions.end();
237  ++it)
238  {
239  if(it->location_number == step.pc->location_number)
240  {
241  from = it;
242  }
243  }
244 
245  INVARIANT(
246  from != from_function.body.instructions.end(),
247  "Failed to find the location number of the loop havoc.");
248 
249  // The violation is caused by the loop havoc
250  // if it is dependent on the loop havoc.
251  if(reachable_set.count(dependence_graph[from].get_node_id()))
252  {
253  result.push_back(
254  loop_idt(step.function_id, original_loop_number_map[step.pc]));
255  return result;
256  }
257  }
258  }
259  return result;
260 }
261 
263  const loop_idt &loop_id,
264  const goto_functiont &function,
265  unsigned location_number_of_target)
266 {
268  loop_id, function, location_number_of_target))
269  {
271  }
272 
274  loop_id, function, location_number_of_target))
275  {
277  }
278 
280 }
281 
283  const loop_idt &loop_id,
284  const goto_functiont &function,
285  unsigned location_number_of_target)
286 {
287  // The transformed loop condition is a set of instructions from
288  // loop havocing instructions
289  // to
290  // if(!guard) GOTO EXIT
291  unsigned location_number_of_havocing = 0;
292  for(auto it = function.body.instructions.begin();
293  it != function.body.instructions.end();
294  ++it)
295  {
296  // Record the location number of the beginning of a transformed loop.
297  if(
298  loop_havoc_set.count(it) &&
299  original_loop_number_map[it] == loop_id.loop_number)
300  {
301  location_number_of_havocing = it->location_number;
302  }
303 
304  // Reach the end of the evaluation of the transformed loop condition.
305  if(location_number_of_havocing != 0 && it->is_goto())
306  {
307  if((location_number_of_havocing < location_number_of_target &&
308  location_number_of_target < it->location_number))
309  {
310  return true;
311  }
312  location_number_of_havocing = 0;
313  }
314  }
315  return false;
316 }
317 
319  const loop_idt &loop_id,
320  const goto_functiont &function,
321  unsigned location_number_of_target)
322 {
323  // The loop body after transformation are instructions from
324  // loop havocing instructions
325  // to
326  // loop end of transformed code.
327  unsigned location_number_of_havocing = 0;
328 
329  for(goto_programt::const_targett it = function.body.instructions.begin();
330  it != function.body.instructions.end();
331  ++it)
332  {
333  // Record the location number of the begin of a transformed loop.
334  if(
335  loop_havoc_set.count(it) &&
336  original_loop_number_map[it] == loop_id.loop_number)
337  {
338  location_number_of_havocing = it->location_number;
339  }
340 
341  // Reach the end of a transformed loop.
342  if(
344  original_loop_number_map[it] == loop_id.loop_number)
345  {
346  INVARIANT(
347  location_number_of_havocing != 0,
348  "We must have entered the transformed loop before reaching the end");
349 
350  // Check if `location_number_of_target` is between the begin and the end
351  // of the transformed loop.
352  if((location_number_of_havocing < location_number_of_target &&
353  location_number_of_target < it->location_number))
354  {
355  return true;
356  }
357  }
358  }
359 
360  return false;
361 }
362 
364  const goto_tracet &goto_trace,
365  const source_locationt &loop_entry_loc)
366 {
367  // Valuations of havoced variables right after havoc instructions.
368  std::unordered_map<exprt, mp_integer, irep_hash> object_sizes;
369  std::unordered_map<exprt, mp_integer, irep_hash> havoced_values;
370  std::unordered_map<exprt, mp_integer, irep_hash> havoced_pointer_offsets;
371 
372  // Loop-entry valuations of havoced variables.
373  std::unordered_map<exprt, mp_integer, irep_hash> loop_entry_values;
374  std::unordered_map<exprt, mp_integer, irep_hash> loop_entry_offsets;
375 
376  // Live variables upon the loop head.
377  std::set<exprt> live_variables;
378 
379  bool entered_loop = false;
380 
381  // Scan the trace step by step to store needed valuations.
382  for(const auto &step : goto_trace.steps)
383  {
384  switch(step.type)
385  {
388  {
389  if(!step.full_lhs_value.is_nil())
390  {
391  // Entered loop?
393  entered_loop = step.full_lhs_value == true_exprt();
394 
395  // skip hidden steps
396  if(step.hidden)
397  break;
398 
399  // Live variables
400  // 1. must be in the same function as the target loop;
401  // 2. alive before entering the target loop;
402  // 3. a pointer or a primitive-typed variable;
403  // TODO: add support for union pointer
404  if(
405  step.pc->source_location().get_function() ==
406  loop_entry_loc.get_function() &&
407  !entered_loop &&
408  (step.full_lhs.type().id() == ID_unsignedbv ||
409  step.full_lhs.type().id() == ID_signedbv ||
410  step.full_lhs.type().id() == ID_pointer) &&
411  step.full_lhs.id() == ID_symbol)
412  {
413  const auto &symbol =
414  expr_try_dynamic_cast<symbol_exprt>(step.full_lhs);
415 
416  // malloc_size is not-hidden tmp variable.
417  if(id2string(symbol->get_identifier()) != "malloc::malloc_size")
418  {
419  live_variables.emplace(step.full_lhs);
420  }
421  }
422 
423  // Record valuation of primitive-typed variable.
424  if(
425  step.full_lhs.type().id() == ID_unsignedbv ||
426  step.full_lhs.type().id() == ID_signedbv)
427  {
428  bool is_signed = step.full_lhs.type().id() == ID_signedbv;
429  const auto &bv_type =
430  type_try_dynamic_cast<bitvector_typet>(step.full_lhs.type());
431  const auto width = bv_type->get_width();
432  // Store the value into the map for loop_entry value if we haven't
433  // entered the loop.
434  if(!entered_loop)
435  {
436  loop_entry_values[step.full_lhs] = bvrep2integer(
437  step.full_lhs_value.get_string(ID_value), width, is_signed);
438  }
439 
440  // Store the value into the the map for havoced value if this step
441  // is a loop havocing instruction.
442  if(loop_havoc_set.count(step.pc))
443  {
444  havoced_values[step.full_lhs] = bvrep2integer(
445  step.full_lhs_value.get_string(ID_value), width, is_signed);
446  }
447  }
448 
449  // Record object_size, pointer_offset, and loop_entry(pointer_offset).
450  if(
452  step.full_lhs_value) &&
454  step.full_lhs_value, SYMEX_DYNAMIC_PREFIX "::dynamic_object"))
455  {
456  const auto &pointer_constant_expr =
457  expr_try_dynamic_cast<annotated_pointer_constant_exprt>(
458  step.full_lhs_value);
459 
460  pointer_arithmetict pointer_arithmetic(
461  pointer_constant_expr->symbolic_pointer());
462  if(pointer_constant_expr->symbolic_pointer().id() == ID_typecast)
463  {
464  pointer_arithmetic = pointer_arithmetict(
465  pointer_constant_expr->symbolic_pointer().operands()[0]);
466  }
467 
468  // Extract object size.
469  exprt &underlying_array = pointer_arithmetic.pointer;
470  // Object sizes are stored in the type of underlying arrays.
471  while(!can_cast_type<array_typet>(underlying_array.type()))
472  {
473  if(
474  underlying_array.id() == ID_address_of ||
475  underlying_array.id() == ID_index ||
476  underlying_array.id() == ID_typecast)
477  {
478  underlying_array = underlying_array.operands()[0];
479  continue;
480  }
481  UNREACHABLE;
482  }
484  pointer_offset_size(to_array_type(underlying_array.type()), ns)
485  .value();
486  object_sizes[step.full_lhs] = object_size;
487 
488  // Extract offsets.
489  mp_integer offset = 0;
490  if(pointer_arithmetic.offset.is_not_nil())
491  {
492  constant_exprt offset_expr =
493  expr_dynamic_cast<constant_exprt>(pointer_arithmetic.offset);
494  offset = bvrep2integer(
495  offset_expr.get_value(), size_type().get_width(), false);
496  }
497 
498  // Store the offset into the map for loop_entry if we haven't
499  // entered the loop.
500  if(!entered_loop)
501  {
502  loop_entry_offsets[step.full_lhs] = offset;
503  }
504 
505  // Store the offset into the the map for havoced offset if this step
506  // is a loop havocing instruction.
507  if(loop_havoc_set.count(step.pc))
508  {
509  havoced_pointer_offsets[step.full_lhs] = offset;
510  }
511  }
512  }
513  }
514 
531  break;
532 
534  UNREACHABLE;
535  }
536  }
537 
538  return cext(
539  object_sizes,
540  havoced_values,
541  havoced_pointer_offsets,
542  loop_entry_values,
543  loop_entry_offsets,
544  live_variables);
545 }
546 
548 {
549  for(const auto &fun_entry : goto_model.goto_functions.function_map)
550  {
551  irep_idt fun_name = fun_entry.first;
552  goto_model.goto_functions.function_map[fun_name].body.swap(
553  original_functions[fun_name]);
554  }
555 }
556 
557 std::optional<cext> cegis_verifiert::verify()
558 {
559  // This method does the following three things to verify the `goto_model` and
560  // return a formatted counterexample if there is any violated property.
561  //
562  // 1. annotate and apply the loop contracts stored in `invariant_candidates`.
563  //
564  // 2. run the CBMC API to verify the instrumented goto model. As the API is
565  // not merged yet, we preprocess the goto model and run the symex checker
566  // on it to simulate CBMC API.
567  // TODO: ^^^ replace the symex checker once the real API is merged.
568  //
569  // 3. construct the formatted counterexample from the violated property and
570  // its trace.
571 
572  // Store the original functions. We will restore them after the verification.
573  for(const auto &fun_entry : goto_model.goto_functions.function_map)
574  {
575  original_functions[fun_entry.first].copy_from(fun_entry.second.body);
576  }
577 
578  // Annotate the candidates to the goto_model for checking.
580 
581  // Annotate assigns
583 
584  // Control verbosity. We allow non-error output message only when verbosity
585  // is set to larger than messaget::M_DEBUG.
586  const unsigned original_verbosity = log.get_message_handler().get_verbosity();
587  if(original_verbosity < messaget::M_DEBUG)
589 
590  // Apply loop contracts we annotated.
592  cont.unwind_transformed_loops = false;
593  cont.apply_loop_contracts();
596 
597  // Get the checker same as CBMC api without any flag.
598  // TODO: replace the checker with CBMC api once it is implemented.
599  ui_message_handlert ui_message_handler(log.get_message_handler());
601  std::unique_ptr<
603  checker = std::make_unique<
605  options, ui_message_handler, goto_model);
606 
607  goto_convert(
611 
612  // Run the checker to get the result.
613  const resultt result = (*checker)();
614 
615  if(original_verbosity >= messaget::M_DEBUG)
616  checker->report();
617 
618  // Restore the verbosity.
619  log.get_message_handler().set_verbosity(original_verbosity);
620 
621  //
622  // Start to construct the counterexample.
623  //
624 
625  if(result == resultt::PASS)
626  {
628  return std::optional<cext>();
629  }
630 
631  if(result == resultt::ERROR || result == resultt::UNKNOWN)
632  {
633  INVARIANT(false, "Verification failed during loop contract synthesis.");
634  }
635 
636  properties = checker->get_properties();
637  auto target_violation = properties.end();
638 
639  // Find target violation---the violation we want to fix next.
640  // A target violation is an assignable violation or the first violation that
641  // is not assignable violation.
642  for(auto it_property = properties.begin(); it_property != properties.end();
643  it_property++)
644  {
645  if(it_property->second.status != property_statust::FAIL)
646  continue;
647 
648  // assignable violation found
649  if(it_property->second.description.find("assignable") != std::string::npos)
650  {
651  target_violation = it_property;
652  break;
653  }
654 
655  // Store the violation that we want to fix with synthesized
656  // assigns/invariant.
657  // ignore ASSERT FALSE
658  if(
659  target_violation == properties.end() &&
660  simplify_expr(it_property->second.pc->condition(), ns) != false_exprt())
661  {
662  target_violation = it_property;
663  }
664  }
665 
666  // All violations are
667  // ASSERT FALSE
668  if(target_violation == properties.end())
669  {
671  return std::optional<cext>();
672  }
673 
674  target_violation_id = target_violation->first;
675 
676  // Decide the violation type from the description of violation
677  cext::violation_typet violation_type =
678  extract_violation_type(target_violation->second.description);
679 
680  // Compute the cause loop---the loop for which we synthesize loop contracts,
681  // and the counterexample.
682 
683  // If the violation is an assignable check, we synthesize assigns targets.
684  // In the case, cause loops are all loops the violation is in. We keep
685  // adding the new assigns target to the most-inner loop that does not
686  // contain the new target until the assignable violation is resolved.
687 
688  // For other cases, we synthesize loop invariant clauses. We synthesize
689  // invariants for one loop at a time. So we return only the first cause loop
690  // although there can be multiple ones.
691 
692  log.debug() << "Start to compute cause loop ids." << messaget::eom;
693  log.debug() << "Violation description: "
694  << target_violation->second.description << messaget::eom;
695 
696  const auto &trace = checker->get_traces()[target_violation->first];
697  // Doing assigns-synthesis or invariant-synthesis
698  if(violation_type == cext::violation_typet::cex_assignable)
699  {
700  cext result(violation_type);
702  result.checked_pointer = static_cast<const exprt &>(
703  target_violation->second.pc->condition().find(ID_checked_assigns));
705  return result;
706  }
707 
708  // We construct the full counterexample only for violations other than
709  // assignable checks.
710 
711  // Although there can be multiple cause loop ids. We only synthesize
712  // loop invariants for the first cause loop.
713  const std::list<loop_idt> cause_loop_ids =
714  get_cause_loop_id(trace, target_violation->second.pc);
715 
716  if(cause_loop_ids.empty())
717  {
718  log.debug() << "No cause loop found!" << messaget::eom;
720 
721  return cext(violation_type);
722  }
723 
724  log.debug() << "Found cause loop with function id: "
725  << cause_loop_ids.front().function_id
726  << ", and loop number: " << cause_loop_ids.front().loop_number
727  << messaget::eom;
728 
729  auto violation_location = cext::violation_locationt::in_loop;
730  // We always strengthen in_clause if the violation is
731  // invariant-not-preserved.
732  if(violation_type != cext::violation_typet::cex_not_preserved)
733  {
734  // Get the location of the violation
735  violation_location = get_violation_location(
736  cause_loop_ids.front(),
737  goto_model.get_goto_function(cause_loop_ids.front().function_id),
738  target_violation->second.pc->location_number);
739  }
740 
742 
743  auto return_cex = build_cex(
744  trace,
746  cause_loop_ids.front().loop_number,
748  .function_map[cause_loop_ids.front().function_id])
749  ->source_location());
750  return_cex.violated_predicate = target_violation->second.pc->condition();
751  return_cex.cause_loop_ids = cause_loop_ids;
752  return_cex.violation_location = violation_location;
753  return_cex.violation_type = violation_type;
754 
755  // The pointer checked in the null-pointer-check violation.
756  if(violation_type == cext::violation_typet::cex_null_pointer)
757  {
758  return_cex.checked_pointer = get_checked_pointer_from_null_pointer_check(
759  target_violation->second.pc->condition());
760  }
761 
762  return return_cex;
763 }
void add_failed_symbols(symbol_table_baset &symbol_table)
Create a failed-dereference symbol for all symbols in the given table that need one (i....
Pointer Dereferencing.
Goto verifier for verifying all properties that stores traces.
void cprover_c_library_factory(const std::set< irep_idt > &functions, const symbol_table_baset &symbol_table, symbol_table_baset &dest_symbol_table, message_handlert &message_handler)
mp_integer bvrep2integer(const irep_idt &src, std::size_t width, bool is_signed)
convert a bit-vector representation (possibly signed) to integer
static bool contains_symbol_prefix(const exprt &expr, const std::string &prefix)
static const exprt & get_checked_pointer_from_null_pointer_check(const exprt &violation)
Verifier for Counterexample-Guided Synthesis.
exprt & lhs()
Definition: std_expr.h:668
exprt & rhs()
Definition: std_expr.h:678
std::optional< cext > verify()
Verify goto_model.
const namespacet ns
irep_idt target_violation_id
void restore_functions()
Restore transformed functions to original functions.
const std::map< loop_idt, std::set< exprt > > & assigns_map
cext build_cex(const goto_tracet &goto_trace, const source_locationt &loop_entry_loc)
cext::violation_locationt get_violation_location(const loop_idt &loop_id, const goto_functiont &function, unsigned location_number_of_target)
std::list< loop_idt > get_cause_loop_id(const goto_tracet &goto_trace, const goto_programt::const_targett violation)
std::unordered_set< goto_programt::const_targett, const_target_hash > loop_havoc_set
Loop havoc instructions instrumented during applying loop contracts.
bool is_instruction_in_transformed_loop(const loop_idt &loop_id, const goto_functiont &function, unsigned location_number_of_target)
Decide whether the target instruction is in the body of the transformed loop specified by loop_id.
const optionst & options
const invariant_mapt & invariant_candidates
std::unordered_map< goto_programt::const_targett, unsigned, const_target_hash > original_loop_number_map
Map from instrumented instructions for loop contracts to their original loop numbers.
cext::violation_typet extract_violation_type(const std::string &description)
goto_modelt & goto_model
std::unordered_map< irep_idt, goto_programt > original_functions
Map from function names to original functions.
void preprocess_goto_model()
Preprocess the goto model to prepare for verification.
propertiest properties
Result counterexample.
std::list< loop_idt > get_cause_loop_id_for_assigns(const goto_tracet &goto_trace)
bool is_instruction_in_transformed_loop_condition(const loop_idt &loop_id, const goto_functiont &function, unsigned location_number_of_target)
Decide whether the target instruction is between the loop-havoc and the evaluation of the loop guard.
Formatted counterexample.
exprt checked_pointer
std::list< loop_idt > cause_loop_ids
violation_typet
violation_locationt
std::unordered_map< goto_programt::const_targett, unsigned, const_target_hash > get_original_loop_number_map() const
Definition: contracts.h:131
void apply_loop_contracts(const std::set< std::string > &to_exclude_from_nondet_init={})
Applies loop contract transformations.
Definition: contracts.cpp:1472
std::unordered_set< goto_programt::const_targett, const_target_hash > get_loop_havoc_set() const
Definition: contracts.h:137
bool unwind_transformed_loops
Definition: contracts.h:145
A constant literal expression.
Definition: std_expr.h:2987
const irep_idt & get_value() const
Definition: std_expr.h:2995
dstringt has one field, an unsigned integer no which is an index into a static table of strings.
Definition: dstring.h:38
bool starts_with(const char *s) const
equivalent of as_string().starts_with(s)
Definition: dstring.h:95
Equality.
Definition: std_expr.h:1361
Base class for all expressions.
Definition: expr.h:56
depth_iteratort depth_end()
Definition: expr.cpp:249
depth_iteratort depth_begin()
Definition: expr.cpp:247
typet & type()
Return the type of the expression.
Definition: expr.h:84
operandst & operands()
Definition: expr.h:94
The Boolean constant false.
Definition: std_expr.h:3064
function_mapt function_map
::goto_functiont goto_functiont
A goto function, consisting of function body (see body) and parameter identifiers (see parameter_iden...
Definition: goto_function.h:24
const goto_functionst::goto_functiont & get_goto_function(const irep_idt &id) override
Get a GOTO function by name, or throw if no such function exists.
Definition: goto_model.h:89
symbol_tablet symbol_table
Symbol table.
Definition: goto_model.h:31
goto_functionst goto_functions
GOTO functions.
Definition: goto_model.h:34
instructionst::const_iterator const_targett
Definition: goto_program.h:615
irep_idt function_id
Definition: goto_trace.h:111
Trace of a GOTO program.
Definition: goto_trace.h:177
goto_trace_stept & get_last_step()
Retrieves the final step in the trace for manipulation (used to fill a trace from code,...
Definition: goto_trace.h:205
stepst steps
Definition: goto_trace.h:180
std::vector< node_indext > get_reachable(node_indext src, bool forwards) const
Run depth-first search on the graph, starting from a single source node.
Definition: graph.h:602
bool is_not_nil() const
Definition: irep.h:368
const irep_idt & id() const
Definition: irep.h:384
void set_verbosity(unsigned _verbosity)
Definition: message.h:53
unsigned get_verbosity() const
Definition: message.h:54
mstreamt & debug() const
Definition: message.h:429
message_handlert & get_message_handler()
Definition: message.h:184
@ M_DEBUG
Definition: message.h:171
@ M_ERROR
Definition: message.h:170
static eomt eom
Definition: message.h:297
A numerical identifier for the object a pointer points to.
const irep_idt & get_function() const
const irep_idt & get_identifier() const
Definition: std_expr.h:160
The Boolean constant true.
Definition: std_expr.h:3055
Verify and use annotated invariants and pre/post-conditions.
void cprover_cpp_library_factory(const std::set< irep_idt > &functions, const symbol_table_baset &symbol_table, symbol_table_baset &dest_symbol_table, message_handlert &message_handler)
Field-Sensitive Program Dependence Analysis, Litvak et al., FSE 2010.
bool is_null_pointer(const constant_exprt &expr)
Returns true if expr has a pointer type and a value NULL; it also returns true when expr has value ze...
Definition: expr_util.cpp:370
void goto_convert(const codet &code, symbol_table_baset &symbol_table, goto_programt &dest, message_handlert &message_handler, const irep_idt &mode)
Goto Programs with Functions.
Utilities for building havoc code for expressions.
Specify write set in function contracts.
const std::string & id2string(const irep_idt &d)
Definition: irep.h:40
Goto Checker using Multi-Path Symbolic Execution.
Options.
bool can_cast_expr< annotated_pointer_constant_exprt >(const exprt &base)
const pointer_object_exprt & to_pointer_object_expr(const exprt &expr)
Cast an exprt to a pointer_object_exprt.
std::optional< mp_integer > pointer_offset_size(const typet &type, const namespacet &ns)
Compute the size of a type in bytes, rounding up to full bytes.
Pointer Logic.
exprt object_size(const exprt &pointer)
Various predicates over pointers in programs.
#define SYMEX_DYNAMIC_PREFIX
bool process_goto_program(goto_modelt &goto_model, const optionst &options, messaget &log)
Common processing and simplification of goto_programts.
@ FAIL
The property was violated.
resultt
The result of goto verifying.
Definition: properties.h:45
@ UNKNOWN
No property was violated, neither was there an error.
@ PASS
No properties were violated.
@ ERROR
An error occurred during goto checking.
void remove_asm(goto_functionst &goto_functions, symbol_tablet &symbol_table, message_handlert &message_handler)
Replaces inline assembly instructions in the goto program (i.e., instructions of kind OTHER with a co...
Definition: remove_asm.cpp:575
bool has_asm(const goto_functionst &goto_functions)
returns true iff the given goto functions use asm instructions
Definition: remove_asm.cpp:598
Remove 'asm' statements by compiling them into suitable standard goto program instructions.
void remove_skip(goto_programt &goto_program, goto_programt::targett begin, goto_programt::targett end)
remove unnecessary skip statements
Definition: remove_skip.cpp:87
Program Transformation.
void label_properties(goto_modelt &goto_model)
Set the properties to check.
exprt simplify_expr(exprt src, const namespacet &ns)
BigInt mp_integer
Definition: smt_terms.h:17
#define UNREACHABLE
This should be used to mark dead code.
Definition: invariant.h:525
const symbol_exprt & to_symbol_expr(const exprt &expr)
Cast an exprt to a symbol_exprt.
Definition: std_expr.h:272
const not_exprt & to_not_expr(const exprt &expr)
Cast an exprt to an not_exprt.
Definition: std_expr.h:2352
bool can_cast_expr< constant_exprt >(const exprt &base)
Definition: std_expr.h:3021
const equal_exprt & to_equal_expr(const exprt &expr)
Cast an exprt to an equal_exprt.
Definition: std_expr.h:1402
bool can_cast_type< array_typet >(const typet &type)
Check whether a reference to a typet is a array_typet.
Definition: std_types.h:875
const array_typet & to_array_type(const typet &type)
Cast a typet to an array_typet.
Definition: std_types.h:888
Loop id used to identify loops.
Definition: loop_ids.h:28
unsigned int loop_number
Definition: loop_ids.h:37
#define size_type
Definition: unistd.c:347
bool is_signed(const typet &t)
Convenience function – is the type signed?
Definition: util.cpp:45
bool is_assignment_to_instrumented_variable(const goto_programt::const_targett &target, std::string var_name)
Return true if target is an assignment to an instrumented variable with name var_name.
Definition: utils.cpp:579
goto_programt::targett get_loop_head(const unsigned int target_loop_number, goto_functiont &function)
Find and return the first instruction of the natural loop with loop_number in function.
Definition: utils.cpp:695
bool is_transformed_loop_head(const goto_programt::const_targett &target)
Return true if target is the head of some transformed loop.
Definition: utils.cpp:563
bool is_transformed_loop_end(const goto_programt::const_targett &target)
Return true if target is the end of some transformed loop.
Definition: utils.cpp:571
void annotate_invariants(const invariant_mapt &invariant_map, goto_modelt &goto_model)
Annotate the invariants in invariant_map to their corresponding loops.
Definition: utils.cpp:700
void annotate_assigns(const std::map< loop_idt, std::set< exprt >> &assigns_map, goto_modelt &goto_model)
Annotate the assigns in assigns_map to their corresponding loops.
Definition: utils.cpp:720
#define ENTERED_LOOP
Definition: utils.h:24