d_collectTermsCache( c ),
d_consTerms( c ),
d_selTerms( c ),
- d_singleton_eq( u )
+ d_singleton_eq( u ),
+ d_lemmas_produced_c( u )
{
// The kinds we are treating as function application in congruence
d_equalityEngine.addFunctionKind(kind::APPLY_CONSTRUCTOR);
if (done() && !fullEffort(e)) {
return;
}
+ Assert( d_pending.empty() && d_pending_merge.empty() );
d_addedLemma = false;
TimerStat::CodeTimer checkTimer(d_checkTime);
assumptions.push_back( eq );
Node lemma = assumptions.size()==1 ? assumptions[0] : NodeManager::currentNM()->mkNode( OR, assumptions );
Trace("dt-singleton") << "*************Singleton equality lemma " << lemma << std::endl;
- d_out->lemma( lemma );
+ doSendLemma( lemma );
}
}
}else{
if( consIndex==-1 ){
consIndex = j;
}
- if( !dt[ j ].isFinite() && ( !options::finiteModelFind() || !dt.isUFinite() ) ) {
+ if( options::finiteModelFind() ? !dt[ j ].isUFinite() : !dt[ j ].isFinite() ) {
if( !eqc || !eqc->d_selectors ){
needSplit = false;
}
NodeBuilder<> nb(kind::OR);
nb << test << test.notNode();
Node lemma = nb;
- d_out->lemma( lemma );
+ doSendLemma( lemma );
d_out->requirePhase( test, true );
}else{
Trace("dt-split") << "*************Split for constructors on " << n << endl;
d_sygus_split->getSygusSplits( n, dt, children, lemmas );
for( unsigned i=0; i<lemmas.size(); i++ ){
Trace("dt-lemma-sygus") << "Dt sygus lemma : " << lemmas[i] << std::endl;
- d_out->lemma( lemmas[i] );
+ doSendLemma( lemmas[i] );
}
}else{
for( unsigned i=0; i<dt.getNumConstructors(); i++ ){
Assert( !children.empty() );
Node lemma = children.size()==1 ? children[0] : NodeManager::currentNM()->mkNode( kind::OR, children );
Trace("dt-split-debug") << "Split lemma is : " << lemma << std::endl;
- d_out->lemma( lemma );
+ doSendLemma( lemma );
}
return;
}
}
}
+ Trace("datatypes-check") << "Finished check effort " << e << std::endl;
if( Debug.isOn("datatypes") || Debug.isOn("datatypes-split") ) {
Notice() << "TheoryDatatypes::check(): done" << endl;
}
lem = Rewriter::rewrite( lem );
}
Trace("dt-lemma") << "Datatypes lemma : " << lem << std::endl;
- d_out->lemma( lem );
+ doSendLemma( lem );
d_addedLemma = true;
}else{
assertFact( fact, exp );
d_pending_merge.clear();
}
+void TheoryDatatypes::doSendLemma( Node lem ) {
+ if( d_lemmas_produced_c.find( lem )==d_lemmas_produced_c.end() ){
+ d_lemmas_produced_c[lem] = true;
+ d_out->lemma( lem );
+ }
+}
+
void TheoryDatatypes::assertFact( Node fact, Node exp ){
Assert( d_pending_merge.empty() );
bool polarity = fact.getKind() != kind::NOT;
d_sygus_sym_break->addTester( atom );
for( unsigned i=0; i<d_sygus_sym_break->d_lemmas.size(); i++ ){
Trace("dt-lemma-sygus") << "Sygus symmetry breaking lemma : " << d_sygus_sym_break->d_lemmas[i] << std::endl;
- d_out->lemma( d_sygus_sym_break->d_lemmas[i] );
+ doSendLemma( d_sygus_sym_break->d_lemmas[i] );
}
d_sygus_sym_break->d_lemmas.clear();
/*
/** called when two equivalance classes have merged */
void TheoryDatatypes::eqNotifyPostMerge(TNode t1, TNode t2){
if( DatatypesRewriter::isTermDatatype( t1 ) ){
+ Debug("datatypes-debug") << "NotifyPostMerge : " << t1 << " " << t2 << std::endl;
d_pending_merge.push_back( t1.eqNode( t2 ) );
}
}
}
EqcInfo* eqc1 = getOrMakeEqcInfo( t1 );
if( eqc1 ){
+ Debug("datatypes-debug") << " merge eqc info " << eqc2 << " into " << eqc1 << std::endl;
if( !eqc1->d_constructor.get().isNull() ){
trep1 = eqc1->d_constructor.get();
}
TNode cons2 = eqc2->d_constructor.get();
//if both have constructor, then either clash or unification
if( !cons1.isNull() && !cons2.isNull() ){
- Debug("datatypes-debug") << "Constructors : " << cons1 << " " << cons2 << std::endl;
+ Debug("datatypes-debug") << " constructors : " << cons1 << " " << cons2 << std::endl;
Node unifEq = cons1.eqNode( cons2 );
/*
std::vector< Node > exp;
if( conf ){
exp.push_back( unifEq );
d_conflictNode = explain( exp );
+ }
*/
std::vector< Node > rew;
if( DatatypesRewriter::checkClash( cons1, cons2, rew ) ){
*/
}
}
- if( !eqc1->d_inst && eqc2->d_inst ){
- eqc1->d_inst = true;
- }
+ Debug("datatypes-debug") << " instantiated : " << eqc1->d_inst << " " << eqc2->d_inst << std::endl;
+ eqc1->d_inst = eqc1->d_inst || eqc2->d_inst;
if( !cons2.isNull() ){
if( cons1.isNull() ){
- Debug("datatypes-debug") << "Must check if it is okay to set the constructor." << std::endl;
+ Debug("datatypes-debug") << " must check if it is okay to set the constructor." << std::endl;
checkInst = true;
addConstructor( eqc2->d_constructor.get(), eqc1, t1 );
if( d_conflict ){
//d_consEqc[t2] = false;
}
}else{
- Debug("datatypes-debug") << "No eqc info for " << t1 << ", must create" << std::endl;
+ Debug("datatypes-debug") << " no eqc info for " << t1 << ", must create" << std::endl;
//just copy the equivalence class information
eqc1 = getOrMakeEqcInfo( t1, true );
eqc1->d_inst.set( eqc2->d_inst );
//merge labels
NodeListMap::iterator lbl_i = d_labels.find( t2 );
if( lbl_i != d_labels.end() ){
- Debug("datatypes-debug") << "Merge labels from " << eqc2 << " " << t2 << std::endl;
+ Debug("datatypes-debug") << " merge labels from " << eqc2 << " " << t2 << std::endl;
NodeList* lbl = (*lbl_i).second;
for( NodeList::const_iterator j = lbl->begin(); j != lbl->end(); ++j ){
addTester( *j, eqc1, t1 );
if( d_conflict ){
- Debug("datatypes-debug") << "Conflict!" << std::endl;
+ Debug("datatypes-debug") << " conflict!" << std::endl;
return;
}
}
}
NodeListMap::iterator sel_i = d_selector_apps.find( t2 );
if( sel_i != d_selector_apps.end() ){
- Debug("datatypes-debug") << "Merge selectors from " << eqc2 << " " << t2 << std::endl;
+ Debug("datatypes-debug") << " merge selectors from " << eqc2 << " " << t2 << std::endl;
NodeList* sel = (*sel_i).second;
for( NodeList::const_iterator j = sel->begin(); j != sel->end(); ++j ){
addSelector( *j, eqc1, t1, eqc2->d_constructor.get().isNull() );
}
}
if( checkInst ){
+ Debug("datatypes-debug") << " checking instantiate" << std::endl;
instantiate( eqc1, t1 );
if( d_conflict ){
return;
}
}
}
+ Debug("datatypes-debug") << "Finished Merge " << t1 << " " << t2 << std::endl;
}
}
if( neqc.isNull() ){
for( unsigned i=0; i<pcons.size(); i++ ){
//must try the infinite ones first
- if( pcons[i] && (r==1)==dt[ i ].isFinite() ){
+ bool cfinite = options::finiteModelFind() ? dt[ i ].isUFinite() : dt[ i ].isFinite();
+ if( pcons[i] && (r==1)==cfinite ){
neqc = DatatypesRewriter::getInstCons( eqc, dt, i );
//for( unsigned j=0; j<neqc.getNumChildren(); j++ ){
// //if( sels[i].find( j )==sels[i].end() && DatatypesRewriter::isTermDatatype( neqc[j] ) ){
Node v2 = NodeManager::currentNM()->mkSkolem( "k2", tn );
a = v1.eqNode( v2 ).negate();
//send out immediately as lemma
- d_out->lemma( a );
+ doSendLemma( a );
Trace("dt-singleton") << "******** assert " << a << " to avoid singleton cardinality for type " << tn << std::endl;
}
d_singleton_lemma[index][tn] = a;
k = NodeManager::currentNM()->mkSkolem( "k", n.getType(), "for dt instantiation" );
Node eq = k.eqNode( n );
Trace("datatypes-infer") << "DtInfer : instantiation ref : " << eq << std::endl;
- d_out->lemma( eq );
+ //doSendLemma( eq );
+ d_pending.push_back( eq );
+ d_pending_exp[ eq ] = d_true;
}
Node n_ic = DatatypesRewriter::getInstCons( k, dt, index );
//Assert( n_ic==Rewriter::rewrite( n_ic ) );
Debug("datatypes-inst") << "DtInstantiate : " << eqc << " " << eq << std::endl;
d_pending.push_back( eq );
d_pending_exp[ eq ] = exp;
+ Trace("datatypes-infer-debug") << "inst : " << eqc << " " << n << std::endl;
Trace("datatypes-infer") << "DtInfer : instantiate : " << eq << " by " << exp << std::endl;
//eqc->d_inst.set( eq );
d_infer.push_back( eq );