CBMC
horn_encoding.cpp
Go to the documentation of this file.
1 /*******************************************************************\
2 
3 Module: Horn-clause Encoding
4 
5 Author: Daniel Kroening
6 
7 Date: June 2015
8 
9 \*******************************************************************/
10 
13 
14 #include "horn_encoding.h"
15 
16 #include <util/c_types.h>
17 #include <util/exception_utils.h>
18 #include <util/format_expr.h>
19 #include <util/mathematical_expr.h>
20 #include <util/pointer_expr.h>
21 #include <util/prefix.h>
22 #include <util/replace_symbol.h>
23 #include <util/std_code.h>
24 #include <util/std_expr.h>
25 
27 
28 #include <solvers/smt2/smt2_conv.h>
29 
30 #include <algorithm>
31 #include <ostream>
32 #include <unordered_set>
33 
34 class state_typet : public typet
35 {
36 public:
37  state_typet() : typet(ID_state)
38  {
39  }
40 };
41 
43 {
45 }
46 
47 static inline symbol_exprt state_expr()
48 {
49  return symbol_exprt(u8"\u03c2", state_typet());
50 }
51 
52 class evaluate_exprt : public binary_exprt
53 {
54 public:
56  : binary_exprt(
57  std::move(state),
58  ID_evaluate,
59  std::move(address),
60  std::move(type))
61  {
62  PRECONDITION(this->state().type().id() == ID_state);
63  PRECONDITION(this->address().type().id() == ID_pointer);
64  }
65 
66  const exprt &state() const
67  {
68  return op0();
69  }
70 
72  {
73  return op0();
74  }
75 
76  const exprt &address() const
77  {
78  return op1();
79  }
80 };
81 
88 inline const evaluate_exprt &to_evaluate_expr(const exprt &expr)
89 {
90  PRECONDITION(expr.id() == ID_evaluate);
91  const evaluate_exprt &ret = static_cast<const evaluate_exprt &>(expr);
92  validate_expr(ret);
93  return ret;
94 }
95 
98 {
99  PRECONDITION(expr.id() == ID_evaluate);
100  evaluate_exprt &ret = static_cast<evaluate_exprt &>(expr);
101  validate_expr(ret);
102  return ret;
103 }
104 
105 class update_state_exprt : public ternary_exprt
106 {
107 public:
109  : ternary_exprt(
110  ID_update_state,
111  std::move(state),
112  std::move(address),
113  std::move(new_value),
114  state_typet())
115  {
116  // PRECONDITION(this->state().id() == ID_state);
117  PRECONDITION(this->address().type().id() == ID_pointer);
118  }
119 
120  const exprt &state() const
121  {
122  return op0();
123  }
124 
126  {
127  return op0();
128  }
129 
130  const exprt &address() const
131  {
132  return op1();
133  }
134 
135  const exprt &new_value() const
136  {
137  return op2();
138  }
139 };
140 
148 {
149  PRECONDITION(expr.id() == ID_update_state);
150  const update_state_exprt &ret = static_cast<const update_state_exprt &>(expr);
151  validate_expr(ret);
152  return ret;
153 }
154 
157 {
158  PRECONDITION(expr.id() == ID_update_state);
159  update_state_exprt &ret = static_cast<update_state_exprt &>(expr);
160  validate_expr(ret);
161  return ret;
162 }
163 
164 class allocate_exprt : public ternary_exprt
165 {
166 public:
168  : ternary_exprt(
169  ID_allocate,
170  std::move(state),
171  std::move(address),
172  std::move(size),
173  state_typet())
174  {
175  PRECONDITION(this->state().type().id() == ID_state);
176  PRECONDITION(this->address().type().id() == ID_pointer);
177  }
178 
179  const exprt &state() const
180  {
181  return op0();
182  }
183 
185  {
186  return op0();
187  }
188 
189  const exprt &address() const
190  {
191  return op1();
192  }
193 
194  const exprt &size() const
195  {
196  return op2();
197  }
198 };
199 
206 inline const allocate_exprt &to_allocate_expr(const exprt &expr)
207 {
208  PRECONDITION(expr.id() == ID_allocate);
209  const allocate_exprt &ret = static_cast<const allocate_exprt &>(expr);
210  validate_expr(ret);
211  return ret;
212 }
213 
214 class encoding_targett
215 {
216 public:
217  virtual void annotation(const std::string &)
218  {
219  }
220 
221  virtual void set_to_true(source_locationt, exprt) = 0;
222 
223  void set_to_true(exprt expr)
224  {
225  set_to_true(source_location, std::move(expr));
226  }
227 
228  void set_source_location(source_locationt __source_location)
229  {
230  source_location = std::move(__source_location);
231  }
232 
233  virtual ~encoding_targett() = default;
234 
235 protected:
237 };
238 
240 {
241 public:
243 
244  using constraintst = std::vector<exprt>;
246 
248  {
251 
252  constraints.emplace_back(std::move(expr));
253  }
254 
255 protected:
257 };
258 
259 static inline void operator<<(encoding_targett &target, exprt constraint)
260 {
261  target.set_to_true(std::move(constraint));
262 }
263 
264 static inline void
266 {
267  for(const auto &constraint : src.constraints)
268  target << constraint;
269 }
270 
272 {
273 public:
274  smt2_encoding_targett(const namespacet &ns, std::ostream &_out)
275  : out(_out),
276  smt2_conv(ns, "", "cprover", "", smt2_convt::solvert::GENERIC, _out)
277  {
279  smt2_conv.use_as_const = true;
280  }
281 
283  {
284  // finish the conversion to SMT-LIB2
285  smt2_conv();
286  }
287 
288  void set_to_true(source_locationt, exprt expr) override
289  {
290  smt2_conv.set_to_true(std::move(expr));
291  }
292 
293  void annotation(const std::string &text) override
294  {
295  if(text.empty())
296  out << '\n';
297  else
298  out << ';' << ' ' << text << '\n';
299  }
300 
301 protected:
302  std::ostream &out;
304 };
305 
307 {
308 public:
309  explicit ascii_encoding_targett(std::ostream &_out) : out(_out)
310  {
311  }
312 
314 
315  void annotation(const std::string &text) override
316  {
317  out << text << '\n';
318  }
319 
320 protected:
321  std::ostream &out;
322  std::size_t counter = 0;
323 };
324 
325 static void find_variables_rec(
326  const exprt &src,
327  std::unordered_set<symbol_exprt, irep_hash> &result)
328 {
329  if(src.id() == ID_object_address)
330  result.insert(to_object_address_expr(src).object_expr());
331  else
332  {
333  for(auto &op : src.operands())
334  find_variables_rec(op, result);
335  }
336 }
337 
338 class state_encodingt
339 {
340 public:
341  explicit state_encodingt(const goto_functionst &__goto_functions)
342  : goto_functions(__goto_functions)
343  {
344  }
345 
347  const goto_functionst::function_mapt::const_iterator,
348  encoding_targett &);
349 
350  void encode(
351  const goto_functiont &,
352  const std::string &state_prefix,
353  const std::string &annotation,
354  const symbol_exprt &entry_state,
355  const exprt &return_lhs,
356  encoding_targett &);
357 
358 protected:
361 
363  symbol_exprt state_expr_with_suffix(loct, const std::string &suffix) const;
364  symbol_exprt out_state_expr(loct, bool taken) const;
366  std::vector<symbol_exprt> incoming_symbols(loct) const;
367  exprt evaluate_expr(loct, const exprt &, const exprt &) const;
369  loct,
370  const exprt &,
371  const exprt &,
372  const std::unordered_set<symbol_exprt, irep_hash> &) const;
374  loct,
375  const exprt &,
376  std::vector<symbol_exprt> &nondet_symbols) const;
377  exprt evaluate_expr(loct, const exprt &) const;
378  exprt address_rec(loct, const exprt &, exprt) const;
386 
387  std::string state_prefix;
388  std::string annotation;
389  loct first_loc;
392  using incomingt =
393  std::map<loct, std::vector<loct>, goto_programt::target_less_than>;
395 };
396 
398 {
399  return state_expr_with_suffix(loc, "");
400 }
401 
403  loct loc,
404  const std::string &suffix) const
405 {
406  irep_idt identifier =
407  state_prefix + std::to_string(loc->location_number) + suffix;
408  return symbol_exprt(identifier, state_predicate_type());
409 }
410 
411 symbol_exprt state_encodingt::out_state_expr(loct loc, bool taken) const
412 {
413  return state_expr_with_suffix(loc, taken ? "T" : "");
414 }
415 
416 std::vector<symbol_exprt> state_encodingt::incoming_symbols(loct loc) const
417 {
418  auto incoming_it = incoming.find(loc);
419 
420  DATA_INVARIANT(incoming_it != incoming.end(), "incoming is complete");
421 
422  std::vector<symbol_exprt> symbols;
423  symbols.reserve(incoming_it->second.size());
424 
425  for(auto &loc_in : incoming_it->second)
426  {
427  std::string suffix;
428 
429  // conditional jump from loc_in to loc?
430  if(
431  loc_in->is_goto() && !loc_in->condition().is_true() &&
432  loc != std::next(loc_in))
433  {
434  suffix = "T";
435  }
436 
437  symbols.push_back(state_expr_with_suffix(loc_in, suffix));
438  }
439 
440  return symbols;
441 }
442 
444 {
445  if(loc == first_loc)
446  return entry_state;
447 
448  auto incoming_symbols = this->incoming_symbols(loc);
449 
450  if(incoming_symbols.size() == 1)
451  return incoming_symbols.front();
452  else
453  return state_expr_with_suffix(loc, "in");
454 }
455 
457  loct loc,
458  const exprt &state,
459  const exprt &what) const
460 {
461  return evaluate_expr_rec(loc, state, what, {});
462 }
463 
465  loct loc,
466  const exprt &what,
467  std::vector<symbol_exprt> &nondet_symbols) const
468 {
469  if(what.id() == ID_side_effect)
470  {
471  auto &side_effect = to_side_effect_expr(what);
472  auto statement = side_effect.get_statement();
473  if(statement == ID_nondet)
474  {
475  irep_idt identifier =
476  "nondet::" + state_prefix + std::to_string(loc->location_number);
477  auto symbol = symbol_exprt(identifier, side_effect.type());
478  nondet_symbols.push_back(symbol);
479  return std::move(symbol);
480  }
481  else
482  return what; // leave it
483  }
484  else
485  {
486  exprt tmp = what;
487  for(auto &op : tmp.operands())
488  op = replace_nondet_rec(loc, op, nondet_symbols);
489  return tmp;
490  }
491 }
492 
494  loct loc,
495  const exprt &state,
496  const exprt &what,
497  const std::unordered_set<symbol_exprt, irep_hash> &bound_symbols) const
498 {
499  PRECONDITION(state.type().id() == ID_state);
500 
501  if(what.id() == ID_symbol)
502  {
503  const auto &symbol_expr = to_symbol_expr(what);
504 
505  if(bound_symbols.find(symbol_expr) == bound_symbols.end())
506  {
507  return evaluate_exprt(state, address_rec(loc, state, what), what.type());
508  }
509  else
510  return what; // leave as is
511  }
512  else if(
513  what.id() == ID_dereference || what.id() == ID_member ||
514  what.id() == ID_index)
515  {
516  return evaluate_exprt(state, address_rec(loc, state, what), what.type());
517  }
518  else if(what.id() == ID_forall || what.id() == ID_exists)
519  {
520  auto new_quantifier_expr = to_quantifier_expr(what); // copy
521  auto new_bound_symbols = bound_symbols; // copy
522 
523  for(const auto &v : new_quantifier_expr.variables())
524  new_bound_symbols.insert(v);
525 
526  new_quantifier_expr.where() = evaluate_expr_rec(
527  loc, state, new_quantifier_expr.where(), new_bound_symbols);
528 
529  return std::move(new_quantifier_expr);
530  }
531  else if(what.id() == ID_address_of)
532  {
533  const auto &object = to_address_of_expr(what).object();
534  return address_rec(loc, state, object);
535  }
536  else if(what.id() == ID_r_ok || what.id() == ID_w_ok || what.id() == ID_rw_ok)
537  {
538  // we need to add the state
539  const auto &r_or_w_ok_expr = to_r_or_w_ok_expr(what);
540  auto pointer =
541  evaluate_expr_rec(loc, state, r_or_w_ok_expr.pointer(), bound_symbols);
542  auto size =
543  evaluate_expr_rec(loc, state, r_or_w_ok_expr.size(), bound_symbols);
544  return ternary_exprt(what.id(), state, pointer, size, what.type());
545  }
546  else
547  {
548  exprt tmp = what;
549  for(auto &op : tmp.operands())
550  op = evaluate_expr_rec(loc, state, op, bound_symbols);
551  return tmp;
552  }
553 }
554 
555 exprt state_encodingt::evaluate_expr(loct loc, const exprt &what) const
556 {
557  return evaluate_expr(loc, in_state_expr(loc), what);
558 }
559 
561 {
562  return lambda_exprt({state_expr()}, std::move(what));
563 }
564 
565 exprt state_encodingt::forall_states_expr(loct loc, exprt what) const
566 {
567  return forall_exprt(
568  {state_expr()},
571  std::move(what)));
572 }
573 
574 exprt state_encodingt::forall_states_expr(loct loc, exprt condition, exprt what)
575  const
576 {
577  return forall_exprt(
578  {state_expr()},
580  and_exprt(
582  std::move(condition)),
583  std::move(what)));
584 }
585 
586 exprt state_encodingt::address_rec(loct loc, const exprt &state, exprt expr)
587  const
588 {
589  if(expr.id() == ID_symbol)
590  {
591  if(expr.type().id() == ID_array)
592  {
593  const auto &element_type = to_array_type(expr.type()).element_type();
594  return object_address_exprt(
595  to_symbol_expr(expr), pointer_type(element_type));
596  }
597  else
598  return object_address_exprt(to_symbol_expr(expr));
599  }
600  else if(expr.id() == ID_member)
601  {
602  auto compound = to_member_expr(expr).struct_op();
603  auto compound_address = address_rec(loc, state, std::move(compound));
604  auto component_name = to_member_expr(expr).get_component_name();
605 
606  if(expr.type().id() == ID_array)
607  {
608  const auto &element_type = to_array_type(expr.type()).element_type();
609  return field_address_exprt(
610  std::move(compound_address),
611  component_name,
612  pointer_type(element_type));
613  }
614  else
615  {
616  return field_address_exprt(
617  std::move(compound_address), component_name, pointer_type(expr.type()));
618  }
619  }
620  else if(expr.id() == ID_index)
621  {
622  auto array = to_index_expr(expr).array();
623  auto index_evaluated =
624  evaluate_expr(loc, state, to_index_expr(expr).index());
625  auto array_address = address_rec(loc, state, std::move(array));
626  return element_address_exprt(
627  std::move(array_address), std::move(index_evaluated));
628  }
629  else if(expr.id() == ID_dereference)
630  return evaluate_expr(loc, state, to_dereference_expr(expr).pointer());
631  else if(expr.id() == ID_string_constant)
632  {
633  // TBD.
635  "can't do string literals", expr.source_location());
636  }
637  else if(expr.id() == ID_array)
638  {
639  // TBD.
641  "can't do array literals", expr.source_location());
642  }
643  else if(expr.id() == ID_union)
644  {
645  // TBD.
647  "can't do union literals", expr.source_location());
648  }
649  else
650  return nil_exprt();
651 }
652 
654  const
655 {
656  auto s = state_expr();
657 
658  auto address = address_rec(loc, s, lhs);
659 
660  exprt rhs_evaluated = evaluate_expr(loc, s, rhs);
661 
662  std::vector<symbol_exprt> nondet_symbols;
663  exprt new_value = replace_nondet_rec(loc, rhs_evaluated, nondet_symbols);
664 
665  auto new_state = update_state_exprt(s, address, new_value);
666 
667  forall_exprt::variablest binding = {state_expr()};
668  binding.insert(binding.end(), nondet_symbols.begin(), nondet_symbols.end());
669 
670  return forall_exprt(
671  std::move(binding),
674  function_application_exprt(out_state_expr(loc), {new_state})));
675 }
676 
677 void state_encodingt::setup_incoming(const goto_functiont &goto_function)
678 {
679  forall_goto_program_instructions(it, goto_function.body)
680  incoming[it];
681 
682  forall_goto_program_instructions(it, goto_function.body)
683  {
684  if(it->is_goto())
685  incoming[it->get_target()].push_back(it);
686  }
687 
688  forall_goto_program_instructions(it, goto_function.body)
689  {
690  auto next = std::next(it);
691  if(it->is_goto() && it->condition().is_true())
692  {
693  }
694  else if(next != goto_function.body.instructions.end())
695  {
696  incoming[next].push_back(it);
697  }
698  }
699 }
700 
702 {
703  if(src.id() == ID_not)
704  return to_not_expr(src).op();
705  else
706  return not_exprt(src);
707 }
708 
711  encoding_targett &dest)
712 {
713  const auto &function = to_symbol_expr(loc->call_function());
714  const auto &type = to_code_type(function.type());
715  auto identifier = function.get_identifier();
716 
717  auto new_annotation = annotation + u8" \u2192 " + id2string(identifier);
718  dest.annotation(new_annotation);
719 
720  // malloc is special-cased
721  if(identifier == "malloc")
722  {
723  auto state = state_expr();
724  PRECONDITION(loc->call_arguments().size() == 1);
725  auto size_evaluated = evaluate_expr(loc, state, loc->call_arguments()[0]);
726 
727  auto lhs_address = address_rec(loc, state, loc->call_lhs());
728  exprt new_state = ternary_exprt(
729  ID_allocate, state, lhs_address, size_evaluated, state_typet());
730  dest << forall_states_expr(
731  loc, function_application_exprt(out_state_expr(loc), {new_state}));
732  return;
733  }
734 
735  // Find the function
736  auto f = goto_functions.function_map.find(identifier);
737  if(f == goto_functions.function_map.end())
738  DATA_INVARIANT(false, "failed find function in function_map");
739 
740  // Do we have a function body?
741  if(!f->second.body_available())
742  {
743  // no function body -- do LHS assignment nondeterministically, if any
744  if(loc->call_lhs().is_not_nil())
745  {
746  auto rhs = side_effect_expr_nondett(
747  loc->call_lhs().type(), loc->source_location());
748  dest << assignment_constraint(loc, loc->call_lhs(), std::move(rhs));
749  }
750  else
751  {
752  // This is a SKIP.
753  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
754  }
755  }
756 
757  // Evaluate the arguments of the call in the 'in state'
758  exprt arguments_state = state_expr();
759 
760  for(std::size_t i = 0; i < type.parameters().size(); i++)
761  {
762  auto address = object_address_exprt(symbol_exprt(
763  f->second.parameter_identifiers[i], type.parameters()[i].type()));
764  auto argument = loc->call_arguments()[i];
765  auto value = evaluate_expr(loc, state_expr(), argument);
766  arguments_state = update_state_exprt(arguments_state, address, value);
767  }
768 
769  // Now assign them
770  auto function_entry_state = state_expr_with_suffix(loc, "Entry");
771  dest << forall_states_expr(
772  loc, function_application_exprt(function_entry_state, {arguments_state}));
773 
774  // now do the body, recursively
775  state_encodingt body_state_encoding(goto_functions);
776  auto new_state_prefix =
777  state_prefix + std::to_string(loc->location_number) + ".";
778  body_state_encoding.encode(
779  f->second,
780  new_state_prefix,
781  new_annotation,
782  function_entry_state,
783  nil_exprt(),
784  dest);
785 
786  // exit state of called function
787  auto exit_loc = std::prev(f->second.body.instructions.end());
788  irep_idt exit_state_identifier =
789  new_state_prefix + std::to_string(exit_loc->location_number);
790  auto exit_state = symbol_exprt(exit_state_identifier, state_predicate_type());
791 
792  // now link up return state
793  dest << equal_exprt(out_state_expr(loc), std::move(exit_state));
794 }
795 
798  encoding_targett &dest)
799 {
800  // Function pointer?
801  const auto &function = loc->call_function();
802  if(function.id() == ID_dereference)
803  {
804  // TBD.
806  "can't do function pointers", loc->source_location());
807  }
808  else if(function.id() == ID_symbol)
809  {
810  function_call_symbol(loc, dest);
811  }
812  else
813  {
815  false, "got function that's neither a symbol nor a function pointer");
816  }
817 }
818 
820  goto_functionst::function_mapt::const_iterator f_entry,
821  encoding_targett &dest)
822 {
823  const auto &goto_function = f_entry->second;
824 
825  if(goto_function.body.instructions.empty())
826  return;
827 
828  // initial state
829  auto in_state = symbol_exprt("SInitial", state_predicate_type());
830 
831  dest << forall_exprt(
832  {state_expr()},
835 
836  auto annotation = id2string(f_entry->first);
837 
838  encode(goto_function, "S", annotation, in_state, nil_exprt(), dest);
839 }
840 
842  const goto_functiont &goto_function,
843  const std::string &state_prefix,
844  const std::string &annotation,
845  const symbol_exprt &entry_state,
846  const exprt &return_lhs,
847  encoding_targett &dest)
848 {
849  first_loc = goto_function.body.instructions.begin();
850  this->state_prefix = state_prefix;
851  this->annotation = annotation;
852  this->entry_state = entry_state;
853  this->return_lhs = return_lhs;
854 
855  setup_incoming(goto_function);
856 
857  // constraints for each instruction
858  forall_goto_program_instructions(loc, goto_function.body)
859  {
860  // pass on the source code location
861  dest.set_source_location(loc->source_location());
862 
863  // constraints on the incoming state
864  {
865  auto incoming_symbols = this->incoming_symbols(loc);
866 
867  if(incoming_symbols.size() >= 2)
868  {
869  auto s = state_expr();
870  for(auto incoming_symbol : incoming_symbols)
871  {
872  dest << forall_exprt(
873  {s},
875  function_application_exprt(incoming_symbol, {s}),
877  }
878  }
879  }
880 
881  if(loc->is_assign())
882  {
883  auto &lhs = loc->assign_lhs();
884  auto &rhs = loc->assign_rhs();
885 
886  if(lhs.type().id() == ID_array)
887  {
888  // skip
889  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
890  }
891  else if(lhs.type().id() == ID_struct_tag)
892  {
893  // skip
894  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
895  }
896  else if(
897  lhs.id() == ID_symbol &&
898  has_prefix(
899  id2string(to_symbol_expr(lhs).get_identifier()), CPROVER_PREFIX) &&
900  to_symbol_expr(lhs).get_identifier() != CPROVER_PREFIX "rounding_mode")
901  {
902  // skip for now
903  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
904  }
905  else
906  dest << assignment_constraint(loc, lhs, rhs);
907  }
908  else if(loc->is_assume())
909  {
910  // we produce ∅ when the assumption is false
911  auto state = state_expr();
912  auto condition_evaluated = evaluate_expr(loc, state, loc->condition());
913 
914  dest << forall_states_expr(
915  loc,
916  condition_evaluated,
918  }
919  else if(loc->is_goto())
920  {
921  // We produce ∅ when the 'other' branch is taken. Get the condition.
922  const auto &condition = loc->condition();
923 
924  if(condition.is_true())
925  {
926  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
927  }
928  else
929  {
930  auto state = state_expr();
931  auto condition_evaluated = evaluate_expr(loc, state, condition);
932 
933  dest << forall_states_expr(
934  loc,
935  condition_evaluated,
936  function_application_exprt(out_state_expr(loc, true), {state}));
937 
938  dest << forall_states_expr(
939  loc,
940  simplifying_not(condition_evaluated),
941  function_application_exprt(out_state_expr(loc, false), {state}));
942  }
943  }
944  else if(loc->is_assert())
945  {
946  // all assertions need to hold
947  dest << forall_states_expr(
948  loc, evaluate_expr(loc, state_expr(), loc->condition()));
949 
950  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
951  }
952  else if(
953  loc->is_skip() || loc->is_assert() || loc->is_location() ||
954  loc->is_end_function() || loc->is_other())
955  {
956  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
957  }
958  else if(loc->is_decl() || loc->is_dead())
959  {
960  // may wish to havoc the symbol
961  dest << equal_exprt(out_state_expr(loc), in_state_expr(loc));
962  }
963  else if(loc->is_function_call())
964  {
965  function_call(loc, dest);
966  }
967  else if(loc->is_set_return_value())
968  {
969  const auto &rhs = loc->return_value();
970 
971  if(return_lhs.is_nil())
972  {
973  // treat these as assignments to a special symbol named 'return_value'
974  auto lhs = symbol_exprt("return_value", rhs.type());
975  dest << assignment_constraint(loc, std::move(lhs), std::move(rhs));
976  }
977  else
978  {
979  }
980  }
981  }
982 }
983 
985  const goto_modelt &goto_model,
986  bool program_is_inlined,
987  encoding_targett &dest)
988 {
989  if(program_is_inlined)
990  {
991  auto f_entry = goto_model.goto_functions.function_map.find(
993 
994  if(f_entry == goto_model.goto_functions.function_map.end())
995  throw incorrect_goto_program_exceptiont("The program has no entry point");
996 
997  dest.annotation("function " + id2string(f_entry->first));
998 
999  state_encodingt{goto_model.goto_functions}(f_entry, dest);
1000  }
1001  else
1002  {
1003  // sort alphabetically
1004  const auto sorted = goto_model.goto_functions.sorted();
1005  const namespacet ns(goto_model.symbol_table);
1006  for(auto &f : sorted)
1007  {
1008  const auto &symbol = ns.lookup(f->first);
1009  if(
1010  f->first == goto_functionst::entry_point() ||
1012  {
1013  dest.annotation("");
1014  dest.annotation("function " + id2string(f->first));
1015  state_encodingt{goto_model.goto_functions}(f, dest);
1016  }
1017  }
1018  }
1019 }
1020 
1021 std::unordered_set<symbol_exprt, irep_hash>
1022 find_variables(const std::vector<exprt> &src)
1023 {
1024  std::unordered_set<symbol_exprt, irep_hash> result;
1025 
1026  for(auto &expr : src)
1027  find_variables_rec(expr, result);
1028 
1029  return result;
1030 }
1031 
1032 static typet
1034 {
1036  for(auto &v : variables)
1037  domain.push_back(v.type());
1038 
1039  return mathematical_function_typet(std::move(domain), bool_typet());
1040 }
1041 
1043 {
1044  // operands first
1045  for(auto &op : src.operands())
1046  op = variable_encoding(op, variables);
1047 
1048  if(src.id() == ID_forall)
1049  {
1050  auto &forall_expr = to_forall_expr(src);
1051  if(
1052  forall_expr.variables().size() >= 1 &&
1053  forall_expr.variables().front().type().id() == ID_state)
1054  {
1055  // replace 'state' by the program variables
1056  forall_exprt::variablest new_variables = variables;
1057  new_variables.insert(
1058  new_variables.end(),
1059  forall_expr.variables().begin() + 1,
1060  forall_expr.variables().end());
1061  forall_expr
1062  .variables() = std::move(new_variables);
1063  return std::move(forall_expr);
1064  }
1065  }
1066  else if(src.id() == ID_function_application)
1067  {
1068  auto &function_application = to_function_application_expr(src);
1069  if(
1070  function_application.arguments().size() == 1 &&
1071  function_application.arguments().front().type().id() == ID_state)
1072  {
1073  function_application.function().type() =
1074  new_state_predicate_type(variables);
1075 
1076  if(function_application.arguments().front().id() == ID_symbol)
1077  {
1078  exprt::operandst new_arguments;
1079  for(auto &v : variables)
1080  new_arguments.push_back(v);
1081  function_application.arguments() = new_arguments;
1082  }
1083  else if(function_application.arguments().front().id() == ID_tuple)
1084  {
1086  function_application.arguments().front().operands().size() ==
1087  variables.size(),
1088  "tuple size must match");
1089  auto operands = function_application.arguments().front().operands();
1090  function_application.arguments() = operands;
1091  }
1092  else
1093  throw analysis_exceptiont("can't convert " + format_to_string(src));
1094  }
1095  else
1096  throw analysis_exceptiont("can't convert " + format_to_string(src));
1097  }
1098  else if(src.id() == ID_evaluate)
1099  {
1100  auto &evaluate_expr = to_evaluate_expr(src);
1101  if(evaluate_expr.address().id() == ID_object_address)
1102  return symbol_exprt(
1103  to_object_address_expr(evaluate_expr.address()).object_expr());
1104  else
1105  throw analysis_exceptiont("can't convert " + format_to_string(src));
1106  }
1107  else if(src.id() == ID_update_state)
1108  {
1109  auto &update_state_expr = to_update_state_expr(src);
1110  if(update_state_expr.address().id() == ID_object_address)
1111  {
1112  auto lhs_symbol =
1113  to_object_address_expr(update_state_expr.address()).object_expr();
1114  exprt::operandst operands;
1115  for(auto &v : variables)
1116  {
1117  if(v == lhs_symbol)
1118  operands.push_back(update_state_expr.new_value());
1119  else
1120  operands.push_back(v);
1121  }
1122  return tuple_exprt(operands, typet(ID_state));
1123  }
1124  else
1125  throw analysis_exceptiont("can't convert " + format_to_string(src));
1126  }
1127 
1128  return src;
1129 }
1130 
1131 void variable_encoding(std::vector<exprt> &constraints)
1132 {
1133  binding_exprt::variablest variables;
1134 
1135  for(auto &v : find_variables(constraints))
1136  variables.push_back(v);
1137 
1138  if(variables.empty())
1139  throw analysis_exceptiont("no variables found");
1140 
1141  // sort alphabetically
1142  std::sort(
1143  variables.begin(),
1144  variables.end(),
1145  [](const symbol_exprt &a, const symbol_exprt &b) {
1146  return id2string(a.get_identifier()) < id2string(b.get_identifier());
1147  });
1148 
1149  for(auto &c : constraints)
1150  c = variable_encoding(c, variables);
1151 }
1152 
1153 void equality_propagation(std::vector<exprt> &constraints)
1154 {
1155  replace_symbolt values;
1156 
1157  std::vector<exprt> new_constraints;
1158  new_constraints.reserve(constraints.size());
1159 
1160  // forward-propagation of equalities
1161  for(auto &expr : constraints)
1162  {
1163  bool retain_constraint = true;
1164 
1165  // apply the map
1166  values(expr);
1167 
1168  if(expr.id() == ID_equal)
1169  {
1170  const auto &equal_expr = to_equal_expr(expr);
1171  if(equal_expr.lhs().id() == ID_symbol)
1172  {
1173  const auto &symbol_expr = to_symbol_expr(equal_expr.lhs());
1174 
1175  // this is a (deliberate) no-op when the symbol is already in the map
1176  if(values.replaces_symbol(symbol_expr.get_identifier()))
1177  {
1178  }
1179  else
1180  {
1181  values.insert(symbol_expr, equal_expr.rhs());
1182  // insertion has happened
1183  retain_constraint = false;
1184  }
1185  }
1186  }
1187 
1188  if(retain_constraint)
1189  new_constraints.push_back(std::move(expr));
1190  }
1191 
1192  constraints = std::move(new_constraints);
1193 
1194  // apply the map again, to catch any backwards definitions
1195  for(auto &expr : constraints)
1196  {
1197  if(expr.id() == ID_equal && to_equal_expr(expr).lhs().id() == ID_symbol)
1198  {
1199  // it's a definition
1200  }
1201  else
1202  {
1203  // apply the map
1204  values(expr);
1205  }
1206  }
1207 }
1208 
1209 void horn_encoding(const goto_modelt &goto_model, std::ostream &out)
1210 {
1211  const namespacet ns(goto_model.symbol_table);
1212 
1213  container_encoding_targett container;
1214  state_encoding(goto_model, true, container);
1215 
1216  equality_propagation(container.constraints);
1217 
1218  variable_encoding(container.constraints);
1219 
1220  smt2_encoding_targett dest(ns, out);
1221  dest << container;
1222 }
uint64_t u8
Definition: bytecode_info.h:58
void validate_expr(const shuffle_vector_exprt &value)
Definition: c_expr.h:81
pointer_typet pointer_type(const typet &subtype)
Definition: c_types.cpp:235
const code_with_contract_typet & to_code_with_contract_type(const typet &type)
Cast a typet to a code_with_contract_typet.
Definition: c_types.h:467
exprt & object()
Definition: pointer_expr.h:549
allocate_exprt(exprt state, exprt address, exprt size)
const exprt & state() const
Definition: state.h:230
const exprt & address() const
const exprt & size() const
Definition: state.h:240
Thrown when an unexpected error occurs during the analysis (e.g., when the SAT solver returns an erro...
Boolean AND.
Definition: std_expr.h:2120
const typet & element_type() const
The type of the elements of the array.
Definition: std_types.h:827
void annotation(const std::string &text) override
ascii_encoding_targett(std::ostream &_out)
void set_to_true(source_locationt, exprt) override
A base class for binary expressions.
Definition: std_expr.h:638
exprt & op1()
Definition: expr.h:136
const exprt & op2() const =delete
exprt & op0()
Definition: expr.h:133
std::vector< symbol_exprt > variablest
Definition: std_expr.h:3101
The Boolean type.
Definition: std_types.h:36
bool has_contract() const
Definition: c_types.h:402
container_encoding_targett()=default
std::vector< exprt > constraintst
void set_to_true(source_locationt source_location, exprt expr) override
void set_to_true(const exprt &)
For a Boolean expression expr, add the constraint 'expr'.
dstringt has one field, an unsigned integer no which is an index into a static table of strings.
Definition: dstring.h:38
Operator to return the address of an array element relative to a base address.
Definition: pointer_expr.h:748
void set_to_true(exprt expr)
virtual void annotation(const std::string &)
virtual ~encoding_targett()=default
void set_source_location(source_locationt __source_location)
virtual void set_to_true(source_locationt, exprt)=0
source_locationt source_location
Equality.
Definition: std_expr.h:1361
const exprt & state() const
Definition: state.h:100
const exprt & address() const
Definition: state.h:110
evaluate_exprt(exprt state, exprt address, typet type)
Base class for all expressions.
Definition: expr.h:56
std::vector< exprt > operandst
Definition: expr.h:58
bool is_true() const
Return whether the expression is a constant representing true.
Definition: expr.cpp:27
source_locationt & add_source_location()
Definition: expr.h:236
const source_locationt & source_location() const
Definition: expr.h:231
typet & type()
Return the type of the expression.
Definition: expr.h:84
operandst & operands()
Definition: expr.h:94
Operator to return the address of a field relative to a base address.
Definition: pointer_expr.h:664
A forall expression.
Application of (mathematical) function.
A collection of goto functions.
function_mapt function_map
std::vector< function_mapt::const_iterator > sorted() const
returns a vector of the iterators in alphabetical order
static irep_idt entry_point()
Get the identifier of the entry point to a goto model.
A goto function, consisting of function body (see body) and parameter identifiers (see parameter_iden...
Definition: goto_function.h:24
goto_programt body
Definition: goto_function.h:26
symbol_tablet symbol_table
Symbol table.
Definition: goto_model.h:31
goto_functionst goto_functions
GOTO functions.
Definition: goto_model.h:34
instructionst instructions
The list of instructions in the goto program.
Definition: goto_program.h:622
instructionst::const_iterator const_targett
Definition: goto_program.h:615
Boolean implication.
Definition: std_expr.h:2183
Thrown when a goto program that's being processed is in an invalid format, for example passing the wr...
exprt & array()
Definition: std_expr.h:1495
bool is_not_nil() const
Definition: irep.h:368
const irep_idt & id() const
Definition: irep.h:384
bool is_nil() const
Definition: irep.h:364
A (mathematical) lambda expression.
A type for mathematical functions (do not confuse with functions/methods in code)
std::vector< typet > domaint
const exprt & struct_op() const
Definition: std_expr.h:2879
irep_idt get_component_name() const
Definition: std_expr.h:2863
A namespacet is essentially one or two symbol tables bound together, to allow for symbol lookups in t...
Definition: namespace.h:94
bool lookup(const irep_idt &name, const symbolt *&symbol) const override
See documentation for namespace_baset::lookup().
Definition: namespace.cpp:148
The NIL expression.
Definition: std_expr.h:3073
Boolean negation.
Definition: std_expr.h:2327
Operator to return the address of an object.
Definition: pointer_expr.h:596
symbol_exprt object_expr() const
Replace a symbol expression by a given expression.
void insert(const class symbol_exprt &old_expr, const exprt &new_expr)
Sets old_expr to be replaced by new_expr if we don't already have a replacement; otherwise does nothi...
bool replaces_symbol(const irep_idt &id) const
A side_effect_exprt that returns a non-deterministically chosen value.
Definition: std_code.h:1520
bool use_as_const
Definition: smt2_conv.h:67
bool use_array_of_bool
Definition: smt2_conv.h:66
void set_to_true(source_locationt, exprt expr) override
smt2_encoding_targett(const namespacet &ns, std::ostream &_out)
void annotation(const std::string &text) override
state_encoding_smt2_convt smt2_conv
static const source_locationt & nil()
static exprt state_lambda_expr(exprt)
exprt evaluate_expr_rec(loct, const exprt &, const exprt &, const std::unordered_set< symbol_exprt, irep_hash > &) const
void setup_incoming(const goto_functiont &)
exprt forall_states_expr(loct, exprt, exprt) const
symbol_exprt in_state_expr(loct) const
void function_call(goto_programt::const_targett, encoding_targett &)
exprt address_rec(loct, const exprt &, exprt) const
goto_programt::const_targett loct
void encode(const goto_functiont &, const irep_idt function_identifier, const std::string &state_prefix, const std::vector< irep_idt > &call_stack, const std::string &annotation, const symbol_exprt &entry_state, const exprt &return_lhs, encoding_targett &)
std::vector< symbol_exprt > incoming_symbols(loct) const
std::string annotation
symbol_exprt out_state_expr(loct) const
state_encodingt(const goto_functionst &__goto_functions)
std::map< loct, std::vector< loct >, goto_programt::target_less_than > incomingt
exprt replace_nondet_rec(loct, const exprt &, std::vector< symbol_exprt > &nondet_symbols) const
void function_call_symbol(goto_programt::const_targett, encoding_targett &)
symbol_exprt entry_state
exprt evaluate_expr(loct, const exprt &, const exprt &) const
symbol_exprt out_state_expr(loct, bool taken) const
symbol_exprt state_expr_with_suffix(loct, const std::string &suffix) const
exprt evaluate_expr(loct, const exprt &) const
exprt assignment_constraint(loct, exprt lhs, exprt rhs) const
std::string state_prefix
std::vector< symbol_exprt > incoming_symbols(loct) const
void operator()(const goto_functionst::function_mapt::const_iterator, encoding_targett &)
static exprt state_lambda_expr(exprt)
exprt forall_states_expr(loct, exprt) const
const goto_functionst & goto_functions
Expression to hold a symbol (variable)
Definition: std_expr.h:131
An expression with three operands.
Definition: std_expr.h:67
exprt & op1()
Definition: expr.h:136
exprt & op2()
Definition: expr.h:139
exprt & op0()
Definition: expr.h:133
The Boolean constant true.
Definition: std_expr.h:3055
The type of an expression, extends irept.
Definition: type.h:29
const exprt & op() const
Definition: std_expr.h:391
const exprt & state() const
Definition: state.h:165
update_state_exprt(exprt state, exprt address, exprt new_value)
const exprt & new_value() const
Definition: state.h:180
const exprt & address() const
Definition: state.h:175
bool has_prefix(const std::string &s, const std::string &prefix)
Definition: converter.cpp:13
#define CPROVER_PREFIX
#define forall_expr(it, expr)
Definition: expr.h:32
std::string format_to_string(const T &o)
Definition: format.h:43
Symbol Table + CFG.
#define forall_goto_program_instructions(it, program)
static typet new_state_predicate_type(const binding_exprt::variablest &variables)
void equality_propagation(std::vector< exprt > &constraints)
const update_state_exprt & to_update_state_expr(const exprt &expr)
Cast an exprt to a update_state_exprt.
static void find_variables_rec(const exprt &src, std::unordered_set< symbol_exprt, irep_hash > &result)
static symbol_exprt state_expr()
std::unordered_set< symbol_exprt, irep_hash > find_variables(const std::vector< exprt > &src)
Returns the set of program variables (as identified by object_address expressions) in the given expre...
const evaluate_exprt & to_evaluate_expr(const exprt &expr)
Cast an exprt to a evaluate_exprt.
void horn_encoding(const goto_modelt &goto_model, std::ostream &out)
exprt variable_encoding(exprt src, const binding_exprt::variablest &variables)
const allocate_exprt & to_allocate_expr(const exprt &expr)
Cast an exprt to a allocate_exprt.
void state_encoding(const goto_modelt &goto_model, bool program_is_inlined, encoding_targett &dest)
static void operator<<(encoding_targett &target, exprt constraint)
static exprt simplifying_not(exprt src)
static mathematical_function_typet state_predicate_type()
Horn-clause Encoding.
const std::string & id2string(const irep_idt &d)
Definition: irep.h:40
API to expression classes for 'mathematical' expressions.
const function_application_exprt & to_function_application_expr(const exprt &expr)
Cast an exprt to a function_application_exprt.
const quantifier_exprt & to_quantifier_expr(const exprt &expr)
Cast an exprt to a quantifier_exprt.
const forall_exprt & to_forall_expr(const exprt &expr)
API to expression classes for Pointers.
const r_or_w_ok_exprt & to_r_or_w_ok_expr(const exprt &expr)
Definition: pointer_expr.h:949
const dereference_exprt & to_dereference_expr(const exprt &expr)
Cast an exprt to a dereference_exprt.
Definition: pointer_expr.h:890
const object_address_exprt & to_object_address_expr(const exprt &expr)
Cast an exprt to an object_address_exprt.
Definition: pointer_expr.h:643
const address_of_exprt & to_address_of_expr(const exprt &expr)
Cast an exprt to an address_of_exprt.
Definition: pointer_expr.h:577
#define DATA_INVARIANT(CONDITION, REASON)
This condition should be used to document that assumptions that are made on goto_functions,...
Definition: invariant.h:534
#define PRECONDITION(CONDITION)
Definition: invariant.h:463
side_effect_exprt & to_side_effect_expr(exprt &expr)
Definition: std_code.h:1506
API to expression classes.
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
const equal_exprt & to_equal_expr(const exprt &expr)
Cast an exprt to an equal_exprt.
Definition: std_expr.h:1402
const member_exprt & to_member_expr(const exprt &expr)
Cast an exprt to a member_exprt.
Definition: std_expr.h:2933
const index_exprt & to_index_expr(const exprt &expr)
Cast an exprt to an index_exprt.
Definition: std_expr.h:1533
const code_typet & to_code_type(const typet &type)
Cast a typet to a code_typet.
Definition: std_types.h:788
const array_typet & to_array_type(const typet &type)
Cast a typet to an array_typet.
Definition: std_types.h:888
std::string to_string(const string_not_contains_constraintt &expr)
Used for debug printing.
A total order over targett and const_targett.
Definition: goto_program.h:392
dstringt irep_idt