CBMC
variable_sensitivity_domain.cpp
Go to the documentation of this file.
1 /*******************************************************************\
2 
3 Module: Abstract Interpretation
4 
5 Author: Martin Brain
6 
7 Date: April 2016
8 
9 \*******************************************************************/
10 
12 
13 #include <util/cprover_prefix.h>
14 #include <util/pointer_expr.h>
15 #include <util/symbol_table.h>
16 
17 #include <algorithm>
18 
19 #ifdef DEBUG
20 # include <iostream>
21 #endif
22 
24  const irep_idt &function_from,
25  trace_ptrt trace_from,
26  const irep_idt &function_to,
27  trace_ptrt trace_to,
28  ai_baset &ai,
29  const namespacet &ns)
30 {
31  locationt from{trace_from->current_location()};
32  locationt to{trace_to->current_location()};
33 
34 #ifdef DEBUG
35  std::cout << "Transform from/to:\n";
36  std::cout << from->location_number << " --> " << to->location_number << '\n';
37 #endif
38 
39  const goto_programt::instructiont &instruction = *from;
40  switch(instruction.type())
41  {
42  case DECL:
43  {
44  abstract_object_pointert top_object =
47  instruction.decl_symbol().type(), ns, true, false)
48  ->write_location_context(from);
49  abstract_state.assign(instruction.decl_symbol(), top_object, ns);
50  }
51  // We now store top.
52  break;
53 
54  case DEAD:
55  // Remove symbol from map, the only time this occurs now (keep TOP.)
56  // It should be the case that DEAD only provides symbols for deletion.
57  abstract_state.erase(instruction.dead_symbol());
58  break;
59 
60  case ASSIGN:
61  {
62  // TODO : check return values
64  abstract_state.eval(instruction.assign_rhs(), ns)
65  ->write_location_context(from);
66  abstract_state.assign(instruction.assign_lhs(), rhs, ns);
67  }
68  break;
69 
70  case GOTO:
71  {
73  {
74  // Get the next line
75  locationt next = from;
76  next++;
77  // Is this a GOTO to the next line (i.e. pointless)
78  if(next != from->get_target())
79  {
80  if(to == from->get_target())
81  {
82  // The AI is exploring the branch where the jump is taken
83  assume(instruction.condition(), ns);
84  }
85  else
86  {
87  // Exploring the path where the jump is not taken - therefore assume
88  // the condition is false
89  assume(not_exprt(instruction.condition()), ns);
90  }
91  }
92  // ignore jumps to the next line, we can assume nothing
93  }
94  }
95  break;
96 
97  case ASSUME:
98  assume(instruction.condition(), ns);
99  break;
100 
101  case FUNCTION_CALL:
102  {
103  transform_function_call(from, to, ai, ns);
104  break;
105  }
106 
107  case END_FUNCTION:
108  {
109  // erase parameters
110 
111  const irep_idt id = function_from;
112  const symbolt &symbol = ns.lookup(id);
113 
114  const code_typet &type = to_code_type(symbol.type);
115 
116  for(const auto &param : type.parameters())
117  {
118  // Top the arguments to the function
120  symbol_exprt(param.get_identifier(), param.type()),
121  abstract_state.abstract_object_factory(param.type(), ns, true, false),
122  ns);
123  }
124 
125  break;
126  }
127 
128  /***************************************************************/
129 
130  case ASSERT:
131  // Conditions on the program, do not alter the data or information
132  // flow and thus can be ignored.
133  // Checking of assertions can only be reasonably done once the fix-point
134  // has been computed, i.e. after all of the calls to transform.
135  break;
136 
137  case SKIP:
138  case LOCATION:
139  // Can ignore
140  break;
141 
142  case SET_RETURN_VALUE:
143  throw "the SET_RETURN_VALUE instructions should be removed first";
144 
145  case START_THREAD:
146  case END_THREAD:
147  case ATOMIC_BEGIN:
148  case ATOMIC_END:
149  throw "threading not supported";
150 
151  case THROW:
152  case CATCH:
153  throw "exceptions not handled";
154 
155  case OTHER:
156  // throw "other";
157  break;
158 
159  case NO_INSTRUCTION_TYPE:
160  break;
161  case INCOMPLETE_GOTO:
162  break;
163  default:
164  throw "unrecognised instruction type";
165  }
166 
167  DATA_INVARIANT(abstract_state.verify(), "Structural invariant");
168 }
169 
171  std::ostream &out,
172  const ai_baset &ai,
173  const namespacet &ns) const
174 {
175  abstract_state.output(out, ai, ns);
176 }
177 
179 {
180  return abstract_state.to_predicate();
181 }
182 
184  const exprt &expr,
185  const namespacet &ns) const
186 {
187  auto result = abstract_state.eval(expr, ns);
188  return result->to_predicate(expr);
189 }
190 
192  const exprt::operandst &exprs,
193  const namespacet &ns) const
194 {
195  if(exprs.empty())
196  return false_exprt();
197  if(exprs.size() == 1)
198  return to_predicate(exprs.front(), ns);
199 
200  auto predicates = std::vector<exprt>{};
202  exprs.cbegin(),
203  exprs.cend(),
204  std::back_inserter(predicates),
205  [this, &ns](const exprt &expr) { return to_predicate(expr, ns); });
206  return and_exprt(predicates);
207 }
208 
210 {
212  return;
213 }
214 
216 {
218 }
219 
221 {
223 }
224 
227  trace_ptrt from,
228  trace_ptrt to)
229 {
230 #ifdef DEBUG
231  std::cout << "Merging from/to:\n "
232  << from->current_location()->location_number << " --> "
233  << to->current_location()->location_number << '\n';
234 #endif
235  auto widen_mode =
236  from->should_widen(*to) ? widen_modet::could_widen : widen_modet::no;
237  // Use the abstract_environment merge
238  bool any_changes =
239  abstract_state.merge(b.abstract_state, to->current_location(), widen_mode);
240 
241  DATA_INVARIANT(abstract_state.verify(), "Structural invariant");
242  return any_changes;
243 }
244 
246  exprt &condition,
247  const namespacet &ns) const
248 {
249  abstract_object_pointert res = abstract_state.eval(condition, ns);
250  exprt c = res->to_constant();
251 
252  if(c.id() == ID_nil)
253  {
254  bool no_simplification = true;
255 
256  // Try to simplify recursively any child operations
257  for(exprt &op : condition.operands())
258  {
259  no_simplification &= ai_simplify(op, ns);
260  }
261 
262  return no_simplification;
263  }
264  else
265  {
266  bool condition_changed = (condition != c);
267  condition = c;
268  return !condition_changed;
269  }
270 }
271 
273 {
274  return abstract_state.is_bottom();
275 }
276 
278 {
279  return abstract_state.is_top();
280 }
281 
283  const variable_sensitivity_domaint &other) const
284 {
287 }
288 
290  locationt from,
291  locationt to,
292  ai_baset &ai,
293  const namespacet &ns)
294 {
295  PRECONDITION(from->type() == FUNCTION_CALL);
296 
297  const exprt &function = from->call_function();
298 
299  const locationt next = std::next(from);
300 
301  if(function.id() == ID_symbol)
302  {
303  // called function identifier
304  const symbol_exprt &symbol_expr = to_symbol_expr(function);
305  const irep_idt function_id = symbol_expr.get_identifier();
306 
307  const code_function_callt::argumentst &called_arguments =
308  from->call_arguments();
309 
310  if(to->location_number == next->location_number)
311  {
312  if(ignore_function_call_transform(function_id))
313  {
314  return;
315  }
316 
317  if(0) // Sound on opaque function calls
318  {
319  abstract_state.havoc("opaque function call");
320  }
321  else
322  {
323  // For any parameter that is a non-const pointer, top the value it is
324  // pointing at.
325  for(const exprt &called_arg : called_arguments)
326  {
327  if(
328  called_arg.type().id() == ID_pointer &&
329  !to_pointer_type(called_arg.type())
330  .base_type()
331  .get_bool(ID_C_constant))
332  {
333  abstract_object_pointert pointer_value =
334  abstract_state.eval(called_arg, ns);
335 
336  CHECK_RETURN(pointer_value);
337 
338  // Write top to the pointer
339  pointer_value->write(
341  ns,
342  std::stack<exprt>(),
343  nil_exprt(),
345  to_pointer_type(called_arg.type()).base_type(),
346  ns,
347  true,
348  false),
349  false);
350  }
351  }
352 
353  // Top any global values
354  for(const auto &symbol : ns.get_symbol_table().symbols)
355  {
356  if(symbol.second.is_static_lifetime)
357  {
359  symbol_exprt(symbol.first, symbol.second.type),
361  symbol.second.type, ns, true, false),
362  ns);
363  }
364  }
365  }
366  }
367  else
368  {
369  // we have an actual call
370  const symbolt &symbol = ns.lookup(function_id);
371  const code_typet &code_type = to_code_type(symbol.type);
372  const code_typet::parameterst &declaration_parameters =
373  code_type.parameters();
374 
375  code_typet::parameterst::const_iterator parameter_it =
376  declaration_parameters.begin();
377 
378  for(const exprt &called_arg : called_arguments)
379  {
380  if(parameter_it == declaration_parameters.end())
381  {
382  INVARIANT(
383  code_type.has_ellipsis(), "Only case for insufficient args");
384  break;
385  }
386 
387  // Evaluate the expression that is being
388  // passed into the function call (called_arg)
389  abstract_object_pointert param_val =
390  abstract_state.eval(called_arg, ns)->write_location_context(from);
391 
392  // Assign the evaluated value to the symbol associated with the
393  // parameter of the function
394  const symbol_exprt parameter_expr(
395  parameter_it->get_identifier(), called_arg.type());
396  abstract_state.assign(parameter_expr, param_val, ns);
397 
398  ++parameter_it;
399  }
400 
401  // Too few arguments so invalid code
403  parameter_it == declaration_parameters.end(),
404  "Number of arguments should match parameters");
405  }
406  }
407  else
408  {
409  PRECONDITION(to == next);
410  abstract_state.havoc("unknown opaque function call");
411  }
412 }
413 
415  const irep_idt &function_id) const
416 {
417  static const std::set<irep_idt> ignored_internal_function = {
418  CPROVER_PREFIX "set_must",
419  CPROVER_PREFIX "get_must",
420  CPROVER_PREFIX "set_may",
421  CPROVER_PREFIX "get_may",
422  CPROVER_PREFIX "cleanup",
423  CPROVER_PREFIX "clear_may",
424  CPROVER_PREFIX "clear_must"};
425 
426  return ignored_internal_function.find(function_id) !=
427  ignored_internal_function.cend();
428 }
429 
431  const ai_domain_baset &function_call,
432  const ai_domain_baset &function_start,
433  const ai_domain_baset &function_end,
434  const namespacet &ns)
435 {
436  const variable_sensitivity_domaint &cast_function_call =
437  static_cast<const variable_sensitivity_domaint &>(function_call);
438 
439  const variable_sensitivity_domaint &cast_function_start =
440  static_cast<const variable_sensitivity_domaint &>(function_start);
441 
442  const variable_sensitivity_domaint &cast_function_end =
443  static_cast<const variable_sensitivity_domaint &>(function_end);
444 
445  const std::vector<irep_idt> &modified_symbol_names =
446  cast_function_start.get_modified_symbols(cast_function_end);
447 
448  std::vector<symbol_exprt> modified_symbols;
449  modified_symbols.reserve(modified_symbol_names.size());
451  modified_symbol_names.begin(),
452  modified_symbol_names.end(),
453  std::back_inserter(modified_symbols),
454  [&ns](const irep_idt &id) { return ns.lookup(id).symbol_expr(); });
455 
456  abstract_state = cast_function_call.abstract_state;
457  apply_domain(modified_symbols, cast_function_end, ns);
458 
459  return;
460 }
461 
463  std::vector<symbol_exprt> modified_symbols,
464  const variable_sensitivity_domaint &source,
465  const namespacet &ns)
466 {
467  for(const auto &symbol : modified_symbols)
468  {
469  abstract_object_pointert value = source.abstract_state.eval(symbol, ns);
470  abstract_state.assign(symbol, value, ns);
471  }
472 }
473 
475 {
476  ai_simplify(expr, ns);
477  abstract_state.assume(expr, ns);
478 }
479 
480 #ifdef ENABLE_STATS
482 variable_sensitivity_domaint::gather_statistics(const namespacet &ns) const
483 {
485 }
486 #endif
dstringt
dstringt has one field, an unsigned integer no which is an index into a static table of strings.
Definition: dstring.h:36
code_typet::has_ellipsis
bool has_ellipsis() const
Definition: std_types.h:611
variable_sensitivity_domaint::to_predicate
exprt to_predicate() const override
Gives a Boolean condition that is true for all values represented by the domain.
Definition: variable_sensitivity_domain.cpp:178
abstract_object_pointert
sharing_ptrt< class abstract_objectt > abstract_object_pointert
Definition: abstract_object.h:69
SET_RETURN_VALUE
@ SET_RETURN_VALUE
Definition: goto_program.h:45
abstract_environmentt::to_predicate
exprt to_predicate() const
Gives a boolean condition that is true for all values represented by the environment.
Definition: abstract_environment.cpp:425
widen_modet::no
@ no
CHECK_RETURN
#define CHECK_RETURN(CONDITION)
Definition: invariant.h:495
code_typet::parameterst
std::vector< parametert > parameterst
Definition: std_types.h:541
transform
static abstract_object_pointert transform(const exprt &expr, const std::vector< abstract_object_pointert > &operands, const abstract_environmentt &environment, const namespacet &ns)
Definition: abstract_value_object.cpp:159
symbolt::type
typet type
Type of symbol.
Definition: symbol.h:31
variable_sensitivity_domaint::is_bottom
bool is_bottom() const override
Find out if the domain is currently unreachable.
Definition: variable_sensitivity_domain.cpp:272
variable_sensitivity_domaint::ignore_function_call_transform
bool ignore_function_call_transform(const irep_idt &function_id) const
Used to specify which CPROVER internal functions should be skipped over when doing function call tran...
Definition: variable_sensitivity_domain.cpp:414
and_exprt
Boolean AND.
Definition: std_expr.h:2070
abstract_environmentt::assign
virtual bool assign(const exprt &expr, const abstract_object_pointert &value, const namespacet &ns)
Assign a value to an expression.
Definition: abstract_environment.cpp:147
abstract_environmentt::verify
bool verify() const
Check the structural invariants are maintained.
Definition: abstract_environment.cpp:449
exprt
Base class for all expressions.
Definition: expr.h:55
abstract_environmentt::output
void output(std::ostream &out, const class ai_baset &ai, const namespacet &ns) const
Print out all the values in the abstract object map.
Definition: abstract_environment.cpp:408
abstract_environmentt::havoc
virtual void havoc(const std::string &havoc_string)
This should be used as a default case / everything else has failed The string is so that I can easily...
Definition: abstract_environment.cpp:379
symbol_exprt
Expression to hold a symbol (variable)
Definition: std_expr.h:112
abstract_environmentt::make_top
void make_top()
Set the domain to top (i.e. everything)
Definition: abstract_environment.cpp:385
goto_programt::instructiont::decl_symbol
const symbol_exprt & decl_symbol() const
Get the declared symbol for DECL.
Definition: goto_program.h:235
ai_domain_baset::trace_ptrt
ai_history_baset::trace_ptrt trace_ptrt
Definition: ai_domain.h:74
variable_sensitivity_domaint::get_modified_symbols
std::vector< irep_idt > get_modified_symbols(const variable_sensitivity_domaint &other) const
Get symbols that have been modified since this domain and other.
Definition: variable_sensitivity_domain.cpp:282
flow_sensitivityt::sensitive
@ sensitive
coverage_criteriont::ASSUME
@ ASSUME
namespacet
A namespacet is essentially one or two symbol tables bound together, to allow for symbol lookups in t...
Definition: namespace.h:90
variable_sensitivity_domaint::make_bottom
void make_bottom() override
Sets the domain to bottom (no states / unreachable).
Definition: variable_sensitivity_domain.cpp:209
THROW
@ THROW
Definition: goto_program.h:50
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
abstract_environmentt::is_top
bool is_top() const
Gets whether the domain is top.
Definition: abstract_environment.cpp:403
coverage_criteriont::LOCATION
@ LOCATION
GOTO
@ GOTO
Definition: goto_program.h:34
to_code_type
const code_typet & to_code_type(const typet &type)
Cast a typet to a code_typet.
Definition: std_types.h:744
goto_programt::instructiont::type
goto_program_instruction_typet type() const
What kind of instruction?
Definition: goto_program.h:351
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
abstract_object_statisticst
Definition: abstract_object_statistics.h:18
goto_programt::instructiont::dead_symbol
const symbol_exprt & dead_symbol() const
Get the symbol for DEAD.
Definition: goto_program.h:251
variable_sensitivity_domaint::apply_domain
void apply_domain(std::vector< symbol_exprt > modified_symbols, const variable_sensitivity_domaint &target, const namespacet &ns)
Given a domain and some symbols, apply those symbols values to the current domain.
Definition: variable_sensitivity_domain.cpp:462
abstract_environmentt::abstract_object_factory
virtual abstract_object_pointert abstract_object_factory(const typet &type, const namespacet &ns, bool top, bool bottom) const
Look at the configuration for the sensitivity and create an appropriate abstract_object.
Definition: abstract_environment.cpp:312
PRECONDITION
#define PRECONDITION(CONDITION)
Definition: invariant.h:463
symbol_exprt::get_identifier
const irep_idt & get_identifier() const
Definition: std_expr.h:142
abstract_environmentt::make_bottom
void make_bottom()
Set the domain to top (i.e. no possible states / unreachable)
Definition: abstract_environment.cpp:392
nil_exprt
The NIL expression.
Definition: std_expr.h:3025
pointer_expr.h
abstract_environmentt::merge
virtual bool merge(const abstract_environmentt &env, const goto_programt::const_targett &merge_location, widen_modet widen_mode)
Computes the join between "this" and "b".
Definition: abstract_environment.cpp:348
variable_sensitivity_domaint::merge_three_way_function_return
virtual void merge_three_way_function_return(const ai_domain_baset &function_call, const ai_domain_baset &function_start, const ai_domain_baset &function_end, const namespacet &ns)
Perform a context aware merge of the changes that have been applied between function_start and the cu...
Definition: variable_sensitivity_domain.cpp:430
variable_sensitivity_domaint::abstract_state
abstract_environmentt abstract_state
Definition: variable_sensitivity_domain.h:266
to_pointer_type
const pointer_typet & to_pointer_type(const typet &type)
Cast a typet to a pointer_typet.
Definition: pointer_expr.h:79
NO_INSTRUCTION_TYPE
@ NO_INSTRUCTION_TYPE
Definition: goto_program.h:33
variable_sensitivity_domaint::transform
void transform(const irep_idt &function_from, trace_ptrt trace_from, const irep_idt &function_to, trace_ptrt trace_to, ai_baset &ai, const namespacet &ns) override
Compute the abstract transformer for a single instruction.
Definition: variable_sensitivity_domain.cpp:23
to_symbol_expr
const symbol_exprt & to_symbol_expr(const exprt &expr)
Cast an exprt to a symbol_exprt.
Definition: std_expr.h:222
code_typet
Base type of functions.
Definition: std_types.h:538
OTHER
@ OTHER
Definition: goto_program.h:37
abstract_environmentt::eval
virtual abstract_object_pointert eval(const exprt &expr, const namespacet &ns) const
These three are really the heart of the method.
Definition: abstract_environment.cpp:94
irept::id
const irep_idt & id() const
Definition: irep.h:396
exprt::operandst
std::vector< exprt > operandst
Definition: expr.h:58
code_function_callt::argumentst
exprt::operandst argumentst
Definition: goto_instruction_code.h:281
false_exprt
The Boolean constant false.
Definition: std_expr.h:3016
variable_sensitivity_domaint::is_top
bool is_top() const override
Is the domain completely top at this state.
Definition: variable_sensitivity_domain.cpp:277
goto_programt::instructiont::assign_rhs
const exprt & assign_rhs() const
Get the rhs of the assignment for ASSIGN.
Definition: goto_program.h:221
END_FUNCTION
@ END_FUNCTION
Definition: goto_program.h:42
SKIP
@ SKIP
Definition: goto_program.h:38
code_typet::parameters
const parameterst & parameters() const
Definition: std_types.h:655
cprover_prefix.h
variable_sensitivity_domaint::output
void output(std::ostream &out, const ai_baset &ai, const namespacet &ns) const override
Basic text output of the abstract domain.
Definition: variable_sensitivity_domain.cpp:170
ASSIGN
@ ASSIGN
Definition: goto_program.h:46
ai_domain_baset::locationt
goto_programt::const_targett locationt
Definition: ai_domain.h:73
variable_sensitivity_domaint::ai_simplify
bool ai_simplify(exprt &condition, const namespacet &ns) const override
Use the information in the domain to simplify the expression with respect to the current location.
Definition: variable_sensitivity_domain.cpp:245
namespacet::get_symbol_table
const symbol_table_baset & get_symbol_table() const
Return first symbol table registered with the namespace.
Definition: namespace.h:123
CATCH
@ CATCH
Definition: goto_program.h:51
DECL
@ DECL
Definition: goto_program.h:47
goto_programt::instructiont::assign_lhs
const exprt & assign_lhs() const
Get the lhs of the assignment for ASSIGN.
Definition: goto_program.h:207
symbolt
Symbol table entry.
Definition: symbol.h:27
variable_sensitivity_domain.h
symbol_table_baset::symbols
const symbolst & symbols
Read-only field, used to look up symbols given their names.
Definition: symbol_table_base.h:30
ai_baset
This is the basic interface of the abstract interpreter with default implementations of the core func...
Definition: ai.h:118
pointer_typet::base_type
const typet & base_type() const
The type of the data what we point to.
Definition: pointer_expr.h:35
variable_sensitivity_domaint
Definition: variable_sensitivity_domain.h:116
variable_sensitivity_domaint::make_entry
void make_entry() override
Set up a reasonable entry-point state.
Definition: variable_sensitivity_domain.cpp:220
CPROVER_PREFIX
#define CPROVER_PREFIX
Definition: cprover_prefix.h:14
START_THREAD
@ START_THREAD
Definition: goto_program.h:39
widen_modet::could_widen
@ could_widen
FUNCTION_CALL
@ FUNCTION_CALL
Definition: goto_program.h:49
abstract_environmentt::modified_symbols
static std::vector< abstract_environmentt::map_keyt > modified_symbols(const abstract_environmentt &first, const abstract_environmentt &second)
For our implementation of variable sensitivity domains, we need to be able to efficiently find symbol...
Definition: abstract_environment.cpp:482
ATOMIC_END
@ ATOMIC_END
Definition: goto_program.h:44
variable_sensitivity_domaint::merge
virtual bool merge(const variable_sensitivity_domaint &b, trace_ptrt from, trace_ptrt to)
Computes the join between "this" and "b".
Definition: variable_sensitivity_domain.cpp:225
goto_programt::instructiont::condition
const exprt & condition() const
Get the condition of gotos, assume, assert.
Definition: goto_program.h:373
DEAD
@ DEAD
Definition: goto_program.h:48
exprt::operands
operandst & operands()
Definition: expr.h:94
ATOMIC_BEGIN
@ ATOMIC_BEGIN
Definition: goto_program.h:43
symbol_table.h
Author: Diffblue Ltd.
ai_domain_baset
The interface offered by a domain, allows code to manipulate domains without knowing their exact type...
Definition: ai_domain.h:54
variable_sensitivity_domaint::assume
void assume(exprt expr, namespacet ns)
Definition: variable_sensitivity_domain.cpp:474
variable_sensitivity_domaint::make_top
void make_top() override
Sets the domain to top (all states).
Definition: variable_sensitivity_domain.cpp:215
ASSERT
@ ASSERT
Definition: goto_program.h:36
variable_sensitivity_domaint::flow_sensitivity
flow_sensitivityt flow_sensitivity
Definition: variable_sensitivity_domain.h:267
goto_programt::instructiont
This class represents an instruction in the GOTO intermediate representation.
Definition: goto_program.h:180
abstract_environmentt::gather_statistics
abstract_object_statisticst gather_statistics(const namespacet &ns) const
Definition: abstract_environment.cpp:527
abstract_environmentt::is_bottom
bool is_bottom() const
Gets whether the domain is bottom.
Definition: abstract_environment.cpp:398
END_THREAD
@ END_THREAD
Definition: goto_program.h:40
INCOMPLETE_GOTO
@ INCOMPLETE_GOTO
Definition: goto_program.h:52
irept::get_bool
bool get_bool(const irep_idt &name) const
Definition: irep.cpp:58
abstract_environmentt::erase
void erase(const symbol_exprt &expr)
Delete a symbol from the map.
Definition: abstract_environment.cpp:476
variable_sensitivity_domaint::transform_function_call
void transform_function_call(locationt from, locationt to, ai_baset &ai, const namespacet &ns)
Used by variable_sensitivity_domaint::transform to handle FUNCTION_CALL transforms.
Definition: variable_sensitivity_domain.cpp:289
validation_modet::INVARIANT
@ INVARIANT
not_exprt
Boolean negation.
Definition: std_expr.h:2277
abstract_environmentt::assume
virtual bool assume(const exprt &expr, const namespacet &ns)
Reduces the domain based on a condition.
Definition: abstract_environment.cpp:257