Expand arith's farkas lemma rule as a macro (#6577)
[cvc5.git] / src / smt / proof_post_processor.cpp
1 /******************************************************************************
2 * Top contributors (to current version):
3 * Andrew Reynolds, Haniel Barbosa, Aina Niemetz
4 *
5 * This file is part of the cvc5 project.
6 *
7 * Copyright (c) 2009-2021 by the authors listed in the file AUTHORS
8 * in the top-level source directory and their institutional affiliations.
9 * All rights reserved. See the file COPYING in the top-level source
10 * directory for licensing information.
11 * ****************************************************************************
12 *
13 * Implementation of module for processing proof nodes.
14 */
15
16 #include "smt/proof_post_processor.h"
17
18 #include "expr/proof_node_manager.h"
19 #include "expr/skolem_manager.h"
20 #include "options/proof_options.h"
21 #include "options/smt_options.h"
22 #include "preprocessing/assertion_pipeline.h"
23 #include "smt/smt_engine.h"
24 #include "smt/smt_statistics_registry.h"
25 #include "theory/builtin/proof_checker.h"
26 #include "theory/rewriter.h"
27 #include "theory/theory.h"
28
29 using namespace cvc5::kind;
30 using namespace cvc5::theory;
31
32 namespace cvc5 {
33 namespace smt {
34
35 ProofPostprocessCallback::ProofPostprocessCallback(ProofNodeManager* pnm,
36 SmtEngine* smte,
37 ProofGenerator* pppg,
38 bool updateScopedAssumptions)
39 : d_pnm(pnm),
40 d_smte(smte),
41 d_pppg(pppg),
42 d_wfpm(pnm),
43 d_updateScopedAssumptions(updateScopedAssumptions)
44 {
45 d_true = NodeManager::currentNM()->mkConst(true);
46 }
47
48 void ProofPostprocessCallback::initializeUpdate()
49 {
50 d_assumpToProof.clear();
51 d_wfAssumptions.clear();
52 }
53
54 void ProofPostprocessCallback::setEliminateRule(PfRule rule)
55 {
56 d_elimRules.insert(rule);
57 }
58
59 bool ProofPostprocessCallback::shouldUpdate(std::shared_ptr<ProofNode> pn,
60 const std::vector<Node>& fa,
61 bool& continueUpdate)
62 {
63 PfRule id = pn->getRule();
64 if (d_elimRules.find(id) != d_elimRules.end())
65 {
66 return true;
67 }
68 // other than elimination rules, we always update assumptions as long as
69 // d_updateScopedAssumptions is true or they are *not* in scope, i.e., not in
70 // fa
71 if (id != PfRule::ASSUME
72 || (!d_updateScopedAssumptions
73 && std::find(fa.begin(), fa.end(), pn->getResult()) != fa.end()))
74 {
75 Trace("smt-proof-pp-debug")
76 << "... not updating in-scope assumption " << pn->getResult() << "\n";
77 return false;
78 }
79 return true;
80 }
81
82 bool ProofPostprocessCallback::update(Node res,
83 PfRule id,
84 const std::vector<Node>& children,
85 const std::vector<Node>& args,
86 CDProof* cdp,
87 bool& continueUpdate)
88 {
89 Trace("smt-proof-pp-debug") << "- Post process " << id << " " << children
90 << " / " << args << std::endl;
91
92 if (id == PfRule::ASSUME)
93 {
94 // we cache based on the assumption node, not the proof node, since there
95 // may be multiple occurrences of the same node.
96 Node f = args[0];
97 std::shared_ptr<ProofNode> pfn;
98 std::map<Node, std::shared_ptr<ProofNode>>::iterator it =
99 d_assumpToProof.find(f);
100 if (it != d_assumpToProof.end())
101 {
102 Trace("smt-proof-pp-debug") << "...already computed" << std::endl;
103 pfn = it->second;
104 }
105 else
106 {
107 Trace("smt-proof-pp-debug") << "...get proof" << std::endl;
108 Assert(d_pppg != nullptr);
109 // get proof from preprocess proof generator
110 pfn = d_pppg->getProofFor(f);
111 Trace("smt-proof-pp-debug") << "...finished get proof" << std::endl;
112 // print for debugging
113 if (pfn == nullptr)
114 {
115 Trace("smt-proof-pp-debug")
116 << "...no proof, possibly an input assumption" << std::endl;
117 }
118 else
119 {
120 Assert(pfn->getResult() == f);
121 if (Trace.isOn("smt-proof-pp"))
122 {
123 Trace("smt-proof-pp")
124 << "=== Connect proof for preprocessing: " << f << std::endl;
125 Trace("smt-proof-pp") << *pfn.get() << std::endl;
126 }
127 }
128 d_assumpToProof[f] = pfn;
129 }
130 if (pfn == nullptr || pfn->getRule() == PfRule::ASSUME)
131 {
132 Trace("smt-proof-pp-debug") << "...do not add proof" << std::endl;
133 // no update
134 return false;
135 }
136 Trace("smt-proof-pp-debug") << "...add proof" << std::endl;
137 // connect the proof
138 cdp->addProof(pfn);
139 return true;
140 }
141 Node ret = expandMacros(id, children, args, cdp);
142 Trace("smt-proof-pp-debug") << "...expanded = " << !ret.isNull() << std::endl;
143 return !ret.isNull();
144 }
145
146 bool ProofPostprocessCallback::updateInternal(Node res,
147 PfRule id,
148 const std::vector<Node>& children,
149 const std::vector<Node>& args,
150 CDProof* cdp)
151 {
152 bool continueUpdate = true;
153 return update(res, id, children, args, cdp, continueUpdate);
154 }
155
156 Node ProofPostprocessCallback::eliminateCrowdingLits(
157 const std::vector<Node>& clauseLits,
158 const std::vector<Node>& targetClauseLits,
159 const std::vector<Node>& children,
160 const std::vector<Node>& args,
161 CDProof* cdp)
162 {
163 Trace("smt-proof-pp-debug2") << push;
164 NodeManager* nm = NodeManager::currentNM();
165 Node trueNode = nm->mkConst(true);
166 // get crowding lits and the position of the last clause that includes
167 // them. The factoring step must be added after the last inclusion and before
168 // its elimination.
169 std::unordered_set<TNode> crowding;
170 std::vector<std::pair<Node, size_t>> lastInclusion;
171 // positions of eliminators of crowding literals, which are the positions of
172 // the clauses that eliminate crowding literals *after* their last inclusion
173 std::vector<size_t> eliminators;
174 for (size_t i = 0, size = clauseLits.size(); i < size; ++i)
175 {
176 if (!crowding.count(clauseLits[i])
177 && std::find(
178 targetClauseLits.begin(), targetClauseLits.end(), clauseLits[i])
179 == targetClauseLits.end())
180 {
181 Node crowdLit = clauseLits[i];
182 crowding.insert(crowdLit);
183 Trace("smt-proof-pp-debug2") << "crowding lit " << crowdLit << "\n";
184 // found crowding lit, now get its last inclusion position, which is the
185 // position of the last resolution link that introduces the crowding
186 // literal. Note that this position has to be *before* the last link, as a
187 // link *after* the last inclusion must eliminate the crowding literal.
188 size_t j;
189 for (j = children.size() - 1; j > 0; --j)
190 {
191 // notice that only non-singleton clauses may be introducing the
192 // crowding literal, so we only care about non-singleton OR nodes. We
193 // check then against the kind and whether the whole OR node occurs as a
194 // pivot of the respective resolution
195 if (children[j - 1].getKind() != kind::OR)
196 {
197 continue;
198 }
199 uint64_t pivotIndex = 2 * (j - 1);
200 if (args[pivotIndex] == children[j - 1]
201 || args[pivotIndex].notNode() == children[j - 1])
202 {
203 continue;
204 }
205 if (std::find(children[j - 1].begin(), children[j - 1].end(), crowdLit)
206 != children[j - 1].end())
207 {
208 break;
209 }
210 }
211 Assert(j > 0);
212 lastInclusion.emplace_back(crowdLit, j - 1);
213
214 Trace("smt-proof-pp-debug2") << "last inc " << j - 1 << "\n";
215 // get elimination position, starting from the following link as the last
216 // inclusion one. The result is the last (in the chain, but first from
217 // this point on) resolution link that eliminates the crowding literal. A
218 // literal l is eliminated by a link if it contains a literal l' with
219 // opposite polarity to l.
220 for (; j < children.size(); ++j)
221 {
222 bool posFirst = args[(2 * j) - 1] == trueNode;
223 Node pivot = args[(2 * j)];
224 Trace("smt-proof-pp-debug2")
225 << "\tcheck w/ args " << posFirst << " / " << pivot << "\n";
226 // To eliminate the crowding literal (crowdLit), the clause must contain
227 // it with opposite polarity. There are three successful cases,
228 // according to the pivot and its sign
229 //
230 // - crowdLit is the same as the pivot and posFirst is true, which means
231 // that the clause contains its negation and eliminates it
232 //
233 // - crowdLit is the negation of the pivot and posFirst is false, so the
234 // clause contains the node whose negation is crowdLit. Note that this
235 // case may either be crowdLit.notNode() == pivot or crowdLit ==
236 // pivot.notNode().
237 if ((crowdLit == pivot && posFirst)
238 || (crowdLit.notNode() == pivot && !posFirst)
239 || (pivot.notNode() == crowdLit && !posFirst))
240 {
241 Trace("smt-proof-pp-debug2") << "\t\tfound it!\n";
242 eliminators.push_back(j);
243 break;
244 }
245 }
246 AlwaysAssert(j < children.size());
247 }
248 }
249 Assert(!lastInclusion.empty());
250 // order map so that we process crowding literals in the order of the clauses
251 // that last introduce them
252 auto cmp = [](std::pair<Node, size_t>& a, std::pair<Node, size_t>& b) {
253 return a.second < b.second;
254 };
255 std::sort(lastInclusion.begin(), lastInclusion.end(), cmp);
256 // order eliminators
257 std::sort(eliminators.begin(), eliminators.end());
258 if (Trace.isOn("smt-proof-pp-debug"))
259 {
260 Trace("smt-proof-pp-debug") << "crowding lits last inclusion:\n";
261 for (const auto& pair : lastInclusion)
262 {
263 Trace("smt-proof-pp-debug")
264 << "\t- [" << pair.second << "] : " << pair.first << "\n";
265 }
266 Trace("smt-proof-pp-debug") << "eliminators:";
267 for (size_t elim : eliminators)
268 {
269 Trace("smt-proof-pp-debug") << " " << elim;
270 }
271 Trace("smt-proof-pp-debug") << "\n";
272 }
273 // TODO (cvc4-wishues/issues/77): implement also simpler version and compare
274 //
275 // We now start to break the chain, one step at a time. Naively this breaking
276 // down would be one resolution/factoring to each crowding literal, but we can
277 // merge some of the cases. Effectively we do the following:
278 //
279 //
280 // lastClause children[start] ... children[end]
281 // ---------------------------------------------- CHAIN_RES
282 // C
283 // ----------- FACTORING
284 // lastClause' children[start'] ... children[end']
285 // -------------------------------------------------------------- CHAIN_RES
286 // ...
287 //
288 // where
289 // lastClause_0 = children[0]
290 // start_0 = 1
291 // end_0 = eliminators[0] - 1
292 // start_i+1 = nextGuardedElimPos - 1
293 //
294 // The important point is how end_i+1 is computed. It is based on what we call
295 // the "nextGuardedElimPos", i.e., the next elimination position that requires
296 // removal of duplicates. The intuition is that a factoring step may eliminate
297 // the duplicates of crowding literals l1 and l2. If the last inclusion of l2
298 // is before the elimination of l1, then we can go ahead and also perform the
299 // elimination of l2 without another factoring. However if another literal l3
300 // has its last inclusion after the elimination of l2, then the elimination of
301 // l3 is the next guarded elimination.
302 //
303 // To do the above computation then we determine, after a resolution/factoring
304 // step, the first crowded literal to have its last inclusion after "end". The
305 // first elimination position to be bigger than the position of that crowded
306 // literal is the next guarded elimination position.
307 size_t lastElim = 0;
308 Node lastClause = children[0];
309 std::vector<Node> childrenRes;
310 std::vector<Node> childrenResArgs;
311 Node resPlaceHolder;
312 size_t nextGuardedElimPos = eliminators[0];
313 do
314 {
315 size_t start = lastElim + 1;
316 size_t end = nextGuardedElimPos - 1;
317 Trace("smt-proof-pp-debug2")
318 << "res with:\n\tlastClause: " << lastClause << "\n\tstart: " << start
319 << "\n\tend: " << end << "\n";
320 childrenRes.push_back(lastClause);
321 // note that the interval of insert is exclusive in the end, so we add 1
322 childrenRes.insert(childrenRes.end(),
323 children.begin() + start,
324 children.begin() + end + 1);
325 childrenResArgs.insert(childrenResArgs.end(),
326 args.begin() + (2 * start) - 1,
327 args.begin() + (2 * end) + 1);
328 Trace("smt-proof-pp-debug2") << "res children: " << childrenRes << "\n";
329 Trace("smt-proof-pp-debug2") << "res args: " << childrenResArgs << "\n";
330 resPlaceHolder = d_pnm->getChecker()->checkDebug(PfRule::CHAIN_RESOLUTION,
331 childrenRes,
332 childrenResArgs,
333 Node::null(),
334 "");
335 Trace("smt-proof-pp-debug2")
336 << "resPlaceHorder: " << resPlaceHolder << "\n";
337 cdp->addStep(
338 resPlaceHolder, PfRule::CHAIN_RESOLUTION, childrenRes, childrenResArgs);
339 // I need to add factoring if end < children.size(). Otherwise, this is
340 // to be handled by the caller
341 if (end < children.size() - 1)
342 {
343 lastClause = d_pnm->getChecker()->checkDebug(
344 PfRule::FACTORING, {resPlaceHolder}, {}, Node::null(), "");
345 if (!lastClause.isNull())
346 {
347 cdp->addStep(lastClause, PfRule::FACTORING, {resPlaceHolder}, {});
348 }
349 else
350 {
351 lastClause = resPlaceHolder;
352 }
353 Trace("smt-proof-pp-debug2") << "lastClause: " << lastClause << "\n";
354 }
355 else
356 {
357 lastClause = resPlaceHolder;
358 break;
359 }
360 // update for next round
361 childrenRes.clear();
362 childrenResArgs.clear();
363 lastElim = end;
364
365 // find the position of the last inclusion of the next crowded literal
366 size_t nextCrowdedInclusionPos = lastInclusion.size();
367 for (size_t i = 0, size = lastInclusion.size(); i < size; ++i)
368 {
369 if (lastInclusion[i].second > lastElim)
370 {
371 nextCrowdedInclusionPos = i;
372 break;
373 }
374 }
375 Trace("smt-proof-pp-debug2")
376 << "nextCrowdedInclusion/Pos: "
377 << lastInclusion[nextCrowdedInclusionPos].second << "/"
378 << nextCrowdedInclusionPos << "\n";
379 // if there is none, then the remaining literals will be used in the next
380 // round
381 if (nextCrowdedInclusionPos == lastInclusion.size())
382 {
383 nextGuardedElimPos = children.size();
384 }
385 else
386 {
387 nextGuardedElimPos = children.size();
388 for (size_t i = 0, size = eliminators.size(); i < size; ++i)
389 {
390 // nextGuardedElimPos is the largest element of
391 // eliminators bigger the next crowded literal's last inclusion
392 if (eliminators[i] > lastInclusion[nextCrowdedInclusionPos].second)
393 {
394 nextGuardedElimPos = eliminators[i];
395 break;
396 }
397 }
398 Assert(nextGuardedElimPos < children.size());
399 }
400 Trace("smt-proof-pp-debug2")
401 << "nextGuardedElimPos: " << nextGuardedElimPos << "\n";
402 } while (true);
403 Trace("smt-proof-pp-debug2") << pop;
404 return lastClause;
405 }
406
407 Node ProofPostprocessCallback::expandMacros(PfRule id,
408 const std::vector<Node>& children,
409 const std::vector<Node>& args,
410 CDProof* cdp)
411 {
412 if (d_elimRules.find(id) == d_elimRules.end())
413 {
414 // not eliminated
415 return Node::null();
416 }
417 // macro elimination
418 if (id == PfRule::MACRO_SR_EQ_INTRO)
419 {
420 // (TRANS
421 // (SUBS <children> :args args[0:1])
422 // (REWRITE :args <t.substitute(x1,t1). ... .substitute(xn,tn)> args[2]))
423 std::vector<Node> tchildren;
424 Node t = args[0];
425 Node ts;
426 if (!children.empty())
427 {
428 std::vector<Node> sargs;
429 sargs.push_back(t);
430 MethodId ids = MethodId::SB_DEFAULT;
431 if (args.size() >= 2)
432 {
433 if (builtin::BuiltinProofRuleChecker::getMethodId(args[1], ids))
434 {
435 sargs.push_back(args[1]);
436 }
437 }
438 MethodId ida = MethodId::SBA_SEQUENTIAL;
439 if (args.size() >= 3)
440 {
441 if (builtin::BuiltinProofRuleChecker::getMethodId(args[2], ida))
442 {
443 sargs.push_back(args[2]);
444 }
445 }
446 ts = builtin::BuiltinProofRuleChecker::applySubstitution(
447 t, children, ids, ida);
448 Trace("smt-proof-pp-debug")
449 << "...eq intro subs equality is " << t << " == " << ts << ", from "
450 << ids << " " << ida << std::endl;
451 if (ts != t)
452 {
453 Node eq = t.eqNode(ts);
454 // apply SUBS proof rule if necessary
455 if (!updateInternal(eq, PfRule::SUBS, children, sargs, cdp))
456 {
457 // if we specified that we did not want to eliminate, add as step
458 cdp->addStep(eq, PfRule::SUBS, children, sargs);
459 }
460 tchildren.push_back(eq);
461 }
462 }
463 else
464 {
465 // no substitute
466 ts = t;
467 }
468 std::vector<Node> rargs;
469 rargs.push_back(ts);
470 MethodId idr = MethodId::RW_REWRITE;
471 if (args.size() >= 4)
472 {
473 if (builtin::BuiltinProofRuleChecker::getMethodId(args[3], idr))
474 {
475 rargs.push_back(args[3]);
476 }
477 }
478 builtin::BuiltinProofRuleChecker* builtinPfC =
479 static_cast<builtin::BuiltinProofRuleChecker*>(
480 d_pnm->getChecker()->getCheckerFor(PfRule::MACRO_SR_EQ_INTRO));
481 Node tr = builtinPfC->applyRewrite(ts, idr);
482 Trace("smt-proof-pp-debug")
483 << "...eq intro rewrite equality is " << ts << " == " << tr << ", from "
484 << idr << std::endl;
485 if (ts != tr)
486 {
487 Node eq = ts.eqNode(tr);
488 // apply REWRITE proof rule
489 if (!updateInternal(eq, PfRule::REWRITE, {}, rargs, cdp))
490 {
491 // if not elimianted, add as step
492 cdp->addStep(eq, PfRule::REWRITE, {}, rargs);
493 }
494 tchildren.push_back(eq);
495 }
496 if (t == tr)
497 {
498 // typically not necessary, but done to be robust
499 cdp->addStep(t.eqNode(tr), PfRule::REFL, {}, {t});
500 return t.eqNode(tr);
501 }
502 // must add TRANS if two step
503 return addProofForTrans(tchildren, cdp);
504 }
505 else if (id == PfRule::MACRO_SR_PRED_INTRO)
506 {
507 std::vector<Node> tchildren;
508 std::vector<Node> sargs = args;
509 // take into account witness form, if necessary
510 bool reqWitness = d_wfpm.requiresWitnessFormIntro(args[0]);
511 Trace("smt-proof-pp-debug")
512 << "...pred intro reqWitness=" << reqWitness << std::endl;
513 // (TRUE_ELIM
514 // (TRANS
515 // (MACRO_SR_EQ_INTRO <children> :args (t args[1:]))
516 // ... proof of apply_SR(t) = toWitness(apply_SR(t)) ...
517 // (MACRO_SR_EQ_INTRO {} {toWitness(apply_SR(t))})
518 // ))
519 // Notice this is an optimized, one sided version of the expansion of
520 // MACRO_SR_PRED_TRANSFORM below.
521 // We call the expandMacros method on MACRO_SR_EQ_INTRO, where notice
522 // that this rule application is immediately expanded in the recursive
523 // call and not added to the proof.
524 Node conc = expandMacros(PfRule::MACRO_SR_EQ_INTRO, children, sargs, cdp);
525 Trace("smt-proof-pp-debug")
526 << "...pred intro conclusion is " << conc << std::endl;
527 Assert(!conc.isNull());
528 Assert(conc.getKind() == EQUAL);
529 Assert(conc[0] == args[0]);
530 tchildren.push_back(conc);
531 if (reqWitness)
532 {
533 Node weq = addProofForWitnessForm(conc[1], cdp);
534 Trace("smt-proof-pp-debug") << "...weq is " << weq << std::endl;
535 if (addToTransChildren(weq, tchildren))
536 {
537 // toWitness(apply_SR(t)) = apply_SR(toWitness(apply_SR(t)))
538 // rewrite again, don't need substitution. Also we always use the
539 // default rewriter, due to the definition of MACRO_SR_PRED_INTRO.
540 Node weqr = expandMacros(PfRule::MACRO_SR_EQ_INTRO, {}, {weq[1]}, cdp);
541 addToTransChildren(weqr, tchildren);
542 }
543 }
544 // apply transitivity if necessary
545 Node eq = addProofForTrans(tchildren, cdp);
546 Assert(!eq.isNull());
547 Assert(eq.getKind() == EQUAL);
548 Assert(eq[0] == args[0]);
549 Assert(eq[1] == d_true);
550
551 cdp->addStep(eq[0], PfRule::TRUE_ELIM, {eq}, {});
552 return eq[0];
553 }
554 else if (id == PfRule::MACRO_SR_PRED_ELIM)
555 {
556 // (EQ_RESOLVE
557 // children[0]
558 // (MACRO_SR_EQ_INTRO children[1:] :args children[0] ++ args))
559 std::vector<Node> schildren(children.begin() + 1, children.end());
560 std::vector<Node> srargs;
561 srargs.push_back(children[0]);
562 srargs.insert(srargs.end(), args.begin(), args.end());
563 Node conc = expandMacros(PfRule::MACRO_SR_EQ_INTRO, schildren, srargs, cdp);
564 Assert(!conc.isNull());
565 Assert(conc.getKind() == EQUAL);
566 Assert(conc[0] == children[0]);
567 // apply equality resolve
568 cdp->addStep(conc[1], PfRule::EQ_RESOLVE, {children[0], conc}, {});
569 return conc[1];
570 }
571 else if (id == PfRule::MACRO_SR_PRED_TRANSFORM)
572 {
573 // (EQ_RESOLVE
574 // children[0]
575 // (TRANS
576 // (MACRO_SR_EQ_INTRO children[1:] :args (children[0] args[1:]))
577 // ... proof of c = wc
578 // (MACRO_SR_EQ_INTRO {} wc)
579 // (SYMM
580 // (MACRO_SR_EQ_INTRO children[1:] :args <args>)
581 // ... proof of a = wa
582 // (MACRO_SR_EQ_INTRO {} wa))))
583 // where
584 // wa = toWitness(apply_SR(args[0])) and
585 // wc = toWitness(apply_SR(children[0])).
586 Trace("smt-proof-pp-debug")
587 << "Transform " << children[0] << " == " << args[0] << std::endl;
588 if (CDProof::isSame(children[0], args[0]))
589 {
590 Trace("smt-proof-pp-debug") << "...nothing to do" << std::endl;
591 // nothing to do
592 return children[0];
593 }
594 std::vector<Node> tchildren;
595 std::vector<Node> schildren(children.begin() + 1, children.end());
596 std::vector<Node> sargs = args;
597 // first, compute if we need
598 bool reqWitness = d_wfpm.requiresWitnessFormTransform(children[0], args[0]);
599 Trace("smt-proof-pp-debug") << "...reqWitness=" << reqWitness << std::endl;
600 // convert both sides, in three steps, take symmetry of second chain
601 for (unsigned r = 0; r < 2; r++)
602 {
603 std::vector<Node> tchildrenr;
604 // first rewrite children[0], then args[0]
605 sargs[0] = r == 0 ? children[0] : args[0];
606 // t = apply_SR(t)
607 Node eq = expandMacros(PfRule::MACRO_SR_EQ_INTRO, schildren, sargs, cdp);
608 Trace("smt-proof-pp-debug")
609 << "transform subs_rewrite (" << r << "): " << eq << std::endl;
610 Assert(!eq.isNull() && eq.getKind() == EQUAL && eq[0] == sargs[0]);
611 addToTransChildren(eq, tchildrenr);
612 // apply_SR(t) = toWitness(apply_SR(t))
613 if (reqWitness)
614 {
615 Node weq = addProofForWitnessForm(eq[1], cdp);
616 Trace("smt-proof-pp-debug")
617 << "transform toWitness (" << r << "): " << weq << std::endl;
618 if (addToTransChildren(weq, tchildrenr))
619 {
620 // toWitness(apply_SR(t)) = apply_SR(toWitness(apply_SR(t)))
621 // rewrite again, don't need substitution. Also, we always use the
622 // default rewriter, due to the definition of MACRO_SR_PRED_TRANSFORM.
623 Node weqr =
624 expandMacros(PfRule::MACRO_SR_EQ_INTRO, {}, {weq[1]}, cdp);
625 Trace("smt-proof-pp-debug") << "transform rewrite_witness (" << r
626 << "): " << weqr << std::endl;
627 addToTransChildren(weqr, tchildrenr);
628 }
629 }
630 Trace("smt-proof-pp-debug")
631 << "transform connect (" << r << ")" << std::endl;
632 // add to overall chain
633 if (r == 0)
634 {
635 // add the current chain to the overall chain
636 tchildren.insert(tchildren.end(), tchildrenr.begin(), tchildrenr.end());
637 }
638 else
639 {
640 // add the current chain to cdp
641 Node eqr = addProofForTrans(tchildrenr, cdp);
642 if (!eqr.isNull())
643 {
644 Trace("smt-proof-pp-debug") << "transform connect sym " << tchildren
645 << " " << eqr << std::endl;
646 // take symmetry of above and add it to the overall chain
647 addToTransChildren(eqr, tchildren, true);
648 }
649 }
650 Trace("smt-proof-pp-debug")
651 << "transform finish (" << r << ")" << std::endl;
652 }
653
654 // apply transitivity if necessary
655 Node eq = addProofForTrans(tchildren, cdp);
656
657 cdp->addStep(args[0], PfRule::EQ_RESOLVE, {children[0], eq}, {});
658 return args[0];
659 }
660 else if (id == PfRule::MACRO_RESOLUTION
661 || id == PfRule::MACRO_RESOLUTION_TRUST)
662 {
663 // first generate the naive chain_resolution
664 std::vector<Node> chainResArgs{args.begin() + 1, args.end()};
665 Node chainConclusion = d_pnm->getChecker()->checkDebug(
666 PfRule::CHAIN_RESOLUTION, children, chainResArgs, Node::null(), "");
667 Trace("smt-proof-pp-debug") << "Original conclusion: " << args[0] << "\n";
668 Trace("smt-proof-pp-debug")
669 << "chainRes conclusion: " << chainConclusion << "\n";
670 // There are n cases:
671 // - if the conclusion is the same, just replace
672 // - if they have the same literals but in different quantity, add a
673 // FACTORING step
674 // - if the order is not the same, add a REORDERING step
675 // - if there are literals in chainConclusion that are not in the original
676 // conclusion, we need to transform the MACRO_RESOLUTION into a series of
677 // CHAIN_RESOLUTION + FACTORING steps, so that we explicitly eliminate all
678 // these "crowding" literals. We do this via FACTORING so we avoid adding
679 // an exponential number of premises, which would happen if we just
680 // repeated in the premises the clauses needed for eliminating crowding
681 // literals, which could themselves add crowding literals.
682 if (chainConclusion == args[0])
683 {
684 cdp->addStep(
685 chainConclusion, PfRule::CHAIN_RESOLUTION, children, chainResArgs);
686 return chainConclusion;
687 }
688 NodeManager* nm = NodeManager::currentNM();
689 // If we got here, then chainConclusion is NECESSARILY an OR node
690 Assert(chainConclusion.getKind() == kind::OR);
691 // get the literals in the chain conclusion
692 std::vector<Node> chainConclusionLits{chainConclusion.begin(),
693 chainConclusion.end()};
694 std::set<Node> chainConclusionLitsSet{chainConclusion.begin(),
695 chainConclusion.end()};
696 // is args[0] a singleton clause? If it's not an OR node, then yes.
697 // Otherwise, it's only a singleton if it occurs in chainConclusionLitsSet
698 std::vector<Node> conclusionLits;
699 // whether conclusion is singleton
700 if (chainConclusionLitsSet.count(args[0]))
701 {
702 conclusionLits.push_back(args[0]);
703 }
704 else
705 {
706 Assert(args[0].getKind() == kind::OR);
707 conclusionLits.insert(
708 conclusionLits.end(), args[0].begin(), args[0].end());
709 }
710 std::set<Node> conclusionLitsSet{conclusionLits.begin(),
711 conclusionLits.end()};
712 // If the sets are different, there are "crowding" literals, i.e. literals
713 // that were removed by implicit multi-usage of premises in the resolution
714 // chain.
715 if (chainConclusionLitsSet != conclusionLitsSet)
716 {
717 chainConclusion = eliminateCrowdingLits(
718 chainConclusionLits, conclusionLits, children, args, cdp);
719 // update vector of lits. Note that the set is no longer used, so we don't
720 // need to update it
721 //
722 // We need again to check whether chainConclusion is a singleton
723 // clause. As above, it's a singleton if it's in the original
724 // chainConclusionLitsSet.
725 chainConclusionLits.clear();
726 if (chainConclusionLitsSet.count(chainConclusion))
727 {
728 chainConclusionLits.push_back(chainConclusion);
729 }
730 else
731 {
732 Assert(chainConclusion.getKind() == kind::OR);
733 chainConclusionLits.insert(chainConclusionLits.end(),
734 chainConclusion.begin(),
735 chainConclusion.end());
736 }
737 }
738 else
739 {
740 cdp->addStep(
741 chainConclusion, PfRule::CHAIN_RESOLUTION, children, chainResArgs);
742 }
743 Trace("smt-proof-pp-debug")
744 << "Conclusion after chain_res/elimCrowd: " << chainConclusion << "\n";
745 Trace("smt-proof-pp-debug")
746 << "Conclusion lits: " << chainConclusionLits << "\n";
747 // Placeholder for running conclusion
748 Node n = chainConclusion;
749 // factoring
750 if (chainConclusionLits.size() != conclusionLits.size())
751 {
752 // We build it rather than taking conclusionLits because the order may be
753 // different
754 std::vector<Node> factoredLits;
755 std::unordered_set<TNode> clauseSet;
756 for (size_t i = 0, size = chainConclusionLits.size(); i < size; ++i)
757 {
758 if (clauseSet.count(chainConclusionLits[i]))
759 {
760 continue;
761 }
762 factoredLits.push_back(n[i]);
763 clauseSet.insert(n[i]);
764 }
765 Node factored = factoredLits.empty()
766 ? nm->mkConst(false)
767 : factoredLits.size() == 1
768 ? factoredLits[0]
769 : nm->mkNode(kind::OR, factoredLits);
770 cdp->addStep(factored, PfRule::FACTORING, {n}, {});
771 n = factored;
772 }
773 // either same node or n as a clause
774 Assert(n == args[0] || n.getKind() == kind::OR);
775 // reordering
776 if (n != args[0])
777 {
778 cdp->addStep(args[0], PfRule::REORDERING, {n}, {args[0]});
779 }
780 return args[0];
781 }
782 else if (id == PfRule::SUBS)
783 {
784 NodeManager* nm = NodeManager::currentNM();
785 // Notice that a naive way to reconstruct SUBS is to do a term conversion
786 // proof for each substitution.
787 // The proof of f(a) * { a -> g(b) } * { b -> c } = f(g(c)) is:
788 // TRANS( CONG{f}( a=g(b) ), CONG{f}( CONG{g}( b=c ) ) )
789 // Notice that more optimal proofs are possible that do a single traversal
790 // over t. This is done by applying later substitutions to the range of
791 // previous substitutions, until a final simultaneous substitution is
792 // applied to t. For instance, in the above example, we first prove:
793 // CONG{g}( b = c )
794 // by applying the second substitution { b -> c } to the range of the first,
795 // giving us a proof of g(b)=g(c). We then construct the updated proof
796 // by tranitivity:
797 // TRANS( a=g(b), CONG{g}( b=c ) )
798 // We then apply the substitution { a -> g(c), b -> c } to f(a), to obtain:
799 // CONG{f}( TRANS( a=g(b), CONG{g}( b=c ) ) )
800 // which notice is more compact than the proof above.
801 Node t = args[0];
802 // get the kind of substitution
803 MethodId ids = MethodId::SB_DEFAULT;
804 if (args.size() >= 2)
805 {
806 builtin::BuiltinProofRuleChecker::getMethodId(args[1], ids);
807 }
808 MethodId ida = MethodId::SBA_SEQUENTIAL;
809 if (args.size() >= 3)
810 {
811 builtin::BuiltinProofRuleChecker::getMethodId(args[2], ida);
812 }
813 std::vector<std::shared_ptr<CDProof>> pfs;
814 std::vector<TNode> vsList;
815 std::vector<TNode> ssList;
816 std::vector<TNode> fromList;
817 std::vector<ProofGenerator*> pgs;
818 // first, compute the entire substitution
819 for (size_t i = 0, nchild = children.size(); i < nchild; i++)
820 {
821 // get the substitution
822 builtin::BuiltinProofRuleChecker::getSubstitutionFor(
823 children[i], vsList, ssList, fromList, ids);
824 // ensure proofs for each formula in fromList
825 if (children[i].getKind() == AND && ids == MethodId::SB_DEFAULT)
826 {
827 for (size_t j = 0, nchildi = children[i].getNumChildren(); j < nchildi;
828 j++)
829 {
830 Node nodej = nm->mkConst(Rational(j));
831 cdp->addStep(
832 children[i][j], PfRule::AND_ELIM, {children[i]}, {nodej});
833 }
834 }
835 }
836 std::vector<Node> vvec;
837 std::vector<Node> svec;
838 for (size_t i = 0, nvs = vsList.size(); i < nvs; i++)
839 {
840 // Note we process in forward order, since later substitution should be
841 // applied to earlier ones, and the last child of a SUBS is processed
842 // first.
843 TNode var = vsList[i];
844 TNode subs = ssList[i];
845 TNode childFrom = fromList[i];
846 Trace("smt-proof-pp-debug")
847 << "...process " << var << " -> " << subs << " (" << childFrom << ", "
848 << ids << ")" << std::endl;
849 // apply the current substitution to the range
850 if (!vvec.empty() && ida == MethodId::SBA_SEQUENTIAL)
851 {
852 Node ss =
853 subs.substitute(vvec.begin(), vvec.end(), svec.begin(), svec.end());
854 if (ss != subs)
855 {
856 Trace("smt-proof-pp-debug")
857 << "......updated to " << var << " -> " << ss
858 << " based on previous substitution" << std::endl;
859 // make the proof for the tranitivity step
860 std::shared_ptr<CDProof> pf = std::make_shared<CDProof>(d_pnm);
861 pfs.push_back(pf);
862 // prove the updated substitution
863 TConvProofGenerator tcg(d_pnm,
864 nullptr,
865 TConvPolicy::ONCE,
866 TConvCachePolicy::NEVER,
867 "nested_SUBS_TConvProofGenerator",
868 nullptr,
869 true);
870 // add previous rewrite steps
871 for (unsigned j = 0, nvars = vvec.size(); j < nvars; j++)
872 {
873 // substitutions are pre-rewrites
874 tcg.addRewriteStep(vvec[j], svec[j], pgs[j], true);
875 }
876 // get the proof for the update to the current substitution
877 Node seqss = subs.eqNode(ss);
878 std::shared_ptr<ProofNode> pfn = tcg.getProofFor(seqss);
879 Assert(pfn != nullptr);
880 // add the proof
881 pf->addProof(pfn);
882 // get proof for childFrom from cdp
883 pfn = cdp->getProofFor(childFrom);
884 pf->addProof(pfn);
885 // ensure we have a proof of var = subs
886 Node veqs = addProofForSubsStep(var, subs, childFrom, pf.get());
887 // transitivity
888 pf->addStep(var.eqNode(ss), PfRule::TRANS, {veqs, seqss}, {});
889 // add to the substitution
890 vvec.push_back(var);
891 svec.push_back(ss);
892 pgs.push_back(pf.get());
893 continue;
894 }
895 }
896 // Just use equality from CDProof, but ensure we have a proof in cdp.
897 // This may involve a TRUE_INTRO/FALSE_INTRO if the substitution step
898 // uses the assumption childFrom as a Boolean assignment (e.g.
899 // childFrom = true if we are using MethodId::SB_LITERAL).
900 addProofForSubsStep(var, subs, childFrom, cdp);
901 vvec.push_back(var);
902 svec.push_back(subs);
903 pgs.push_back(cdp);
904 }
905 // should be implied by the substitution now
906 TConvPolicy tcpolicy = ida == MethodId::SBA_FIXPOINT ? TConvPolicy::FIXPOINT
907 : TConvPolicy::ONCE;
908 TConvProofGenerator tcpg(d_pnm,
909 nullptr,
910 tcpolicy,
911 TConvCachePolicy::NEVER,
912 "SUBS_TConvProofGenerator",
913 nullptr,
914 true);
915 for (unsigned j = 0, nvars = vvec.size(); j < nvars; j++)
916 {
917 // substitutions are pre-rewrites
918 tcpg.addRewriteStep(vvec[j], svec[j], pgs[j], true);
919 }
920 // add the proof constructed by the term conversion utility
921 std::shared_ptr<ProofNode> pfn = tcpg.getProofForRewriting(t);
922 Node eq = pfn->getResult();
923 Node ts = builtin::BuiltinProofRuleChecker::applySubstitution(
924 t, children, ids, ida);
925 Node eqq = t.eqNode(ts);
926 if (eq != eqq)
927 {
928 pfn = nullptr;
929 }
930 // should give a proof, if not, then tcpg does not agree with the
931 // substitution.
932 Assert(pfn != nullptr);
933 if (pfn == nullptr)
934 {
935 AlwaysAssert(false) << "resort to TRUST_SUBS" << std::endl
936 << eq << std::endl
937 << eqq << std::endl
938 << "from " << children << " applied to " << t;
939 cdp->addStep(eqq, PfRule::TRUST_SUBS, {}, {eqq});
940 }
941 else
942 {
943 cdp->addProof(pfn);
944 }
945 return eqq;
946 }
947 else if (id == PfRule::REWRITE)
948 {
949 // get the kind of rewrite
950 MethodId idr = MethodId::RW_REWRITE;
951 if (args.size() >= 2)
952 {
953 builtin::BuiltinProofRuleChecker::getMethodId(args[1], idr);
954 }
955 builtin::BuiltinProofRuleChecker* builtinPfC =
956 static_cast<builtin::BuiltinProofRuleChecker*>(
957 d_pnm->getChecker()->getCheckerFor(PfRule::REWRITE));
958 Node ret = builtinPfC->applyRewrite(args[0], idr);
959 Node eq = args[0].eqNode(ret);
960 if (idr == MethodId::RW_REWRITE || idr == MethodId::RW_REWRITE_EQ_EXT)
961 {
962 // rewrites from theory::Rewriter
963 bool isExtEq = (idr == MethodId::RW_REWRITE_EQ_EXT);
964 // use rewrite with proof interface
965 Rewriter* rr = d_smte->getRewriter();
966 TrustNode trn = rr->rewriteWithProof(args[0], isExtEq);
967 std::shared_ptr<ProofNode> pfn = trn.toProofNode();
968 if (pfn == nullptr)
969 {
970 Trace("smt-proof-pp-debug")
971 << "Use TRUST_REWRITE for " << eq << std::endl;
972 // did not have a proof of rewriting, probably isExtEq is true
973 if (isExtEq)
974 {
975 // update to THEORY_REWRITE with idr
976 Assert(args.size() >= 1);
977 TheoryId theoryId = Theory::theoryOf(args[0].getType());
978 Node tid = builtin::BuiltinProofRuleChecker::mkTheoryIdNode(theoryId);
979 cdp->addStep(eq, PfRule::THEORY_REWRITE, {}, {eq, tid, args[1]});
980 }
981 else
982 {
983 // this should never be applied
984 cdp->addStep(eq, PfRule::TRUST_REWRITE, {}, {eq});
985 }
986 }
987 else
988 {
989 cdp->addProof(pfn);
990 }
991 Assert(trn.getNode() == ret)
992 << "Unexpected rewrite " << args[0] << std::endl
993 << "Got: " << trn.getNode() << std::endl
994 << "Expected: " << ret;
995 }
996 else if (idr == MethodId::RW_EVALUATE)
997 {
998 // change to evaluate, which is never eliminated
999 cdp->addStep(eq, PfRule::EVALUATE, {}, {args[0]});
1000 }
1001 else
1002 {
1003 // don't know how to eliminate
1004 return Node::null();
1005 }
1006 if (args[0] == ret)
1007 {
1008 // should not be necessary typically
1009 cdp->addStep(eq, PfRule::REFL, {}, {args[0]});
1010 }
1011 return eq;
1012 }
1013 else if (id == PfRule::THEORY_REWRITE)
1014 {
1015 Assert(!args.empty());
1016 Node eq = args[0];
1017 Assert(eq.getKind() == EQUAL);
1018 // try to replay theory rewrite
1019 // first, check that maybe its just an evaluation step
1020 ProofChecker* pc = d_pnm->getChecker();
1021 Node ceval =
1022 pc->checkDebug(PfRule::EVALUATE, {}, {eq[0]}, eq, "smt-proof-pp-debug");
1023 if (!ceval.isNull() && ceval == eq)
1024 {
1025 cdp->addStep(eq, PfRule::EVALUATE, {}, {eq[0]});
1026 return eq;
1027 }
1028 // otherwise no update
1029 Trace("final-pf-hole") << "hole: " << id << " : " << eq << std::endl;
1030 }
1031 else if (id == PfRule::MACRO_ARITH_SCALE_SUM_UB)
1032 {
1033 Debug("macro::arith") << "Expand MACRO_ARITH_SCALE_SUM_UB" << std::endl;
1034 if (Debug.isOn("macro::arith"))
1035 {
1036 for (const auto& child : children)
1037 {
1038 Debug("macro::arith") << " child: " << child << std::endl;
1039 }
1040 Debug("macro::arith") << " args: " << args << std::endl;
1041 }
1042 Assert(args.size() == children.size());
1043 NodeManager* nm = NodeManager::currentNM();
1044 ProofStepBuffer steps{d_pnm->getChecker()};
1045
1046 // Scale all children, accumulating
1047 std::vector<Node> scaledRels;
1048 for (size_t i = 0; i < children.size(); ++i)
1049 {
1050 TNode child = children[i];
1051 TNode scalar = args[i];
1052 bool isPos = scalar.getConst<Rational>() > 0;
1053 Node scalarCmp =
1054 nm->mkNode(isPos ? GT : LT, scalar, nm->mkConst(Rational(0)));
1055 // (= scalarCmp true)
1056 Node scalarCmpOrTrue = steps.tryStep(PfRule::EVALUATE, {}, {scalarCmp});
1057 Assert(!scalarCmpOrTrue.isNull());
1058 // scalarCmp
1059 steps.addStep(PfRule::TRUE_ELIM, {scalarCmpOrTrue}, {}, scalarCmp);
1060 // (and scalarCmp relation)
1061 Node scalarCmpAndRel =
1062 steps.tryStep(PfRule::AND_INTRO, {scalarCmp, child}, {});
1063 Assert(!scalarCmpAndRel.isNull());
1064 // (=> (and scalarCmp relation) scaled)
1065 Node impl =
1066 steps.tryStep(isPos ? PfRule::ARITH_MULT_POS : PfRule::ARITH_MULT_NEG,
1067 {},
1068 {scalar, child});
1069 Assert(!impl.isNull());
1070 // scaled
1071 Node scaled =
1072 steps.tryStep(PfRule::MODUS_PONENS, {scalarCmpAndRel, impl}, {});
1073 Assert(!scaled.isNull());
1074 scaledRels.emplace_back(scaled);
1075 }
1076
1077 Node sumBounds = steps.tryStep(PfRule::ARITH_SUM_UB, scaledRels, {});
1078 cdp->addSteps(steps);
1079 Debug("macro::arith") << "Expansion done. Proved: " << sumBounds
1080 << std::endl;
1081 return sumBounds;
1082 }
1083
1084 // TRUST, PREPROCESS, THEORY_LEMMA, THEORY_PREPROCESS?
1085
1086 return Node::null();
1087 }
1088
1089 Node ProofPostprocessCallback::addProofForWitnessForm(Node t, CDProof* cdp)
1090 {
1091 Node tw = SkolemManager::getOriginalForm(t);
1092 Node eq = t.eqNode(tw);
1093 if (t == tw)
1094 {
1095 // not necessary, add REFL step
1096 cdp->addStep(eq, PfRule::REFL, {}, {t});
1097 return eq;
1098 }
1099 std::shared_ptr<ProofNode> pn = d_wfpm.getProofFor(eq);
1100 if (pn != nullptr)
1101 {
1102 // add the proof
1103 cdp->addProof(pn);
1104 }
1105 else
1106 {
1107 Assert(false) << "ProofPostprocessCallback::addProofForWitnessForm: failed "
1108 "to add proof for witness form of "
1109 << t;
1110 }
1111 return eq;
1112 }
1113
1114 Node ProofPostprocessCallback::addProofForTrans(
1115 const std::vector<Node>& tchildren, CDProof* cdp)
1116 {
1117 size_t tsize = tchildren.size();
1118 if (tsize > 1)
1119 {
1120 Node lhs = tchildren[0][0];
1121 Node rhs = tchildren[tsize - 1][1];
1122 Node eq = lhs.eqNode(rhs);
1123 cdp->addStep(eq, PfRule::TRANS, tchildren, {});
1124 return eq;
1125 }
1126 else if (tsize == 1)
1127 {
1128 return tchildren[0];
1129 }
1130 return Node::null();
1131 }
1132
1133 Node ProofPostprocessCallback::addProofForSubsStep(Node var,
1134 Node subs,
1135 Node assump,
1136 CDProof* cdp)
1137 {
1138 // ensure we have a proof of var = subs
1139 Node veqs = var.eqNode(subs);
1140 if (veqs != assump)
1141 {
1142 // should be true intro or false intro
1143 Assert(subs.isConst());
1144 cdp->addStep(
1145 veqs,
1146 subs.getConst<bool>() ? PfRule::TRUE_INTRO : PfRule::FALSE_INTRO,
1147 {assump},
1148 {});
1149 }
1150 return veqs;
1151 }
1152
1153 bool ProofPostprocessCallback::addToTransChildren(Node eq,
1154 std::vector<Node>& tchildren,
1155 bool isSymm)
1156 {
1157 Assert(!eq.isNull());
1158 Assert(eq.getKind() == kind::EQUAL);
1159 if (eq[0] == eq[1])
1160 {
1161 return false;
1162 }
1163 Node equ = isSymm ? eq[1].eqNode(eq[0]) : eq;
1164 Assert(tchildren.empty()
1165 || (tchildren[tchildren.size() - 1].getKind() == kind::EQUAL
1166 && tchildren[tchildren.size() - 1][1] == equ[0]));
1167 tchildren.push_back(equ);
1168 return true;
1169 }
1170
1171 ProofPostprocessFinalCallback::ProofPostprocessFinalCallback(
1172 ProofNodeManager* pnm)
1173 : d_ruleCount(smtStatisticsRegistry().registerHistogram<PfRule>(
1174 "finalProof::ruleCount")),
1175 d_totalRuleCount(
1176 smtStatisticsRegistry().registerInt("finalProof::totalRuleCount")),
1177 d_minPedanticLevel(
1178 smtStatisticsRegistry().registerInt("finalProof::minPedanticLevel")),
1179 d_numFinalProofs(
1180 smtStatisticsRegistry().registerInt("finalProofs::numFinalProofs")),
1181 d_pnm(pnm),
1182 d_pedanticFailure(false)
1183 {
1184 d_minPedanticLevel += 10;
1185 }
1186
1187 void ProofPostprocessFinalCallback::initializeUpdate()
1188 {
1189 d_pedanticFailure = false;
1190 d_pedanticFailureOut.str("");
1191 ++d_numFinalProofs;
1192 }
1193
1194 bool ProofPostprocessFinalCallback::shouldUpdate(std::shared_ptr<ProofNode> pn,
1195 const std::vector<Node>& fa,
1196 bool& continueUpdate)
1197 {
1198 PfRule r = pn->getRule();
1199 // if not doing eager pedantic checking, fail if below threshold
1200 if (!options::proofEagerChecking())
1201 {
1202 if (!d_pedanticFailure)
1203 {
1204 Assert(d_pedanticFailureOut.str().empty());
1205 if (d_pnm->getChecker()->isPedanticFailure(r, d_pedanticFailureOut))
1206 {
1207 d_pedanticFailure = true;
1208 }
1209 }
1210 }
1211 uint32_t plevel = d_pnm->getChecker()->getPedanticLevel(r);
1212 if (plevel != 0)
1213 {
1214 d_minPedanticLevel.minAssign(plevel);
1215 }
1216 // record stats for the rule
1217 d_ruleCount << r;
1218 ++d_totalRuleCount;
1219 return false;
1220 }
1221
1222 bool ProofPostprocessFinalCallback::wasPedanticFailure(std::ostream& out) const
1223 {
1224 if (d_pedanticFailure)
1225 {
1226 out << d_pedanticFailureOut.str();
1227 return true;
1228 }
1229 return false;
1230 }
1231
1232 ProofPostproccess::ProofPostproccess(ProofNodeManager* pnm,
1233 SmtEngine* smte,
1234 ProofGenerator* pppg,
1235 bool updateScopedAssumptions)
1236 : d_pnm(pnm),
1237 d_cb(pnm, smte, pppg, updateScopedAssumptions),
1238 // the update merges subproofs
1239 d_updater(d_pnm, d_cb, true),
1240 d_finalCb(pnm),
1241 d_finalizer(d_pnm, d_finalCb)
1242 {
1243 }
1244
1245 ProofPostproccess::~ProofPostproccess() {}
1246
1247 void ProofPostproccess::process(std::shared_ptr<ProofNode> pf)
1248 {
1249 // Initialize the callback, which computes necessary static information about
1250 // how to process, including how to process assumptions in pf.
1251 d_cb.initializeUpdate();
1252 // now, process
1253 d_updater.process(pf);
1254 // take stats and check pedantic
1255 d_finalCb.initializeUpdate();
1256 d_finalizer.process(pf);
1257
1258 std::stringstream serr;
1259 bool wasPedanticFailure = d_finalCb.wasPedanticFailure(serr);
1260 if (wasPedanticFailure)
1261 {
1262 AlwaysAssert(!wasPedanticFailure)
1263 << "ProofPostproccess::process: pedantic failure:" << std::endl
1264 << serr.str();
1265 }
1266 }
1267
1268 void ProofPostproccess::setEliminateRule(PfRule rule)
1269 {
1270 d_cb.setEliminateRule(rule);
1271 }
1272
1273 } // namespace smt
1274 } // namespace cvc5