mirror of
				https://github.com/Z3Prover/z3
				synced 2025-11-04 05:19:11 +00:00 
			
		
		
		
	add rule unfolding transformation
Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com>
This commit is contained in:
		
							parent
							
								
									2c24f25050
								
							
						
					
					
						commit
						d16db63e56
					
				
					 13 changed files with 160 additions and 138 deletions
				
			
		| 
						 | 
					@ -118,7 +118,7 @@ namespace datalog {
 | 
				
			||||||
         expr_ref level_query = mk_level_predicate(m_query_pred, level);
 | 
					         expr_ref level_query = mk_level_predicate(m_query_pred, level);
 | 
				
			||||||
         model_ref md;
 | 
					         model_ref md;
 | 
				
			||||||
         proof_ref pr(m);
 | 
					         proof_ref pr(m);
 | 
				
			||||||
         rule_unifier unifier(rm, m_ctx, m);
 | 
					         rule_unifier unifier(m_ctx);
 | 
				
			||||||
         m_solver.get_model(md);
 | 
					         m_solver.get_model(md);
 | 
				
			||||||
         func_decl* pred = m_query_pred;
 | 
					         func_decl* pred = m_query_pred;
 | 
				
			||||||
         SASSERT(m.is_true(md->get_const_interp(to_app(level_query)->get_decl())));
 | 
					         SASSERT(m.is_true(md->get_const_interp(to_app(level_query)->get_decl())));
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -91,7 +91,7 @@ namespace datalog {
 | 
				
			||||||
    }
 | 
					    }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    void rule_unifier::apply(
 | 
					    void rule_unifier::apply(
 | 
				
			||||||
        rule& r, bool is_tgt, unsigned skipped_index, 
 | 
					        rule const& r, bool is_tgt, unsigned skipped_index, 
 | 
				
			||||||
        app_ref_vector& res, svector<bool>& res_neg) {
 | 
					        app_ref_vector& res, svector<bool>& res_neg) {
 | 
				
			||||||
        unsigned rule_len = r.get_tail_size();
 | 
					        unsigned rule_len = r.get_tail_size();
 | 
				
			||||||
        for (unsigned i = 0; i < rule_len; i++) {
 | 
					        for (unsigned i = 0; i < rule_len; i++) {
 | 
				
			||||||
| 
						 | 
					@ -104,7 +104,7 @@ namespace datalog {
 | 
				
			||||||
        }
 | 
					        }
 | 
				
			||||||
    }
 | 
					    }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    bool rule_unifier::apply(rule& tgt, unsigned tail_index, rule& src, rule_ref& res) {
 | 
					    bool rule_unifier::apply(rule const& tgt, unsigned tail_index, rule const& src, rule_ref& res) {
 | 
				
			||||||
        SASSERT(m_ready);
 | 
					        SASSERT(m_ready);
 | 
				
			||||||
        app_ref new_head(m);
 | 
					        app_ref new_head(m);
 | 
				
			||||||
        app_ref_vector tail(m);
 | 
					        app_ref_vector tail(m);
 | 
				
			||||||
| 
						 | 
					@ -116,7 +116,7 @@ namespace datalog {
 | 
				
			||||||
        mk_rule_inliner::remove_duplicate_tails(tail, tail_neg);
 | 
					        mk_rule_inliner::remove_duplicate_tails(tail, tail_neg);
 | 
				
			||||||
        SASSERT(tail.size()==tail_neg.size());
 | 
					        SASSERT(tail.size()==tail_neg.size());
 | 
				
			||||||
        res = m_rm.mk(new_head, tail.size(), tail.c_ptr(), tail_neg.c_ptr());
 | 
					        res = m_rm.mk(new_head, tail.size(), tail.c_ptr(), tail_neg.c_ptr());
 | 
				
			||||||
        res->set_accounting_parent_object(m_context, &tgt);
 | 
					        res->set_accounting_parent_object(m_context, const_cast<rule*>(&tgt));
 | 
				
			||||||
        res->norm_vars(m_rm);
 | 
					        res->norm_vars(m_rm);
 | 
				
			||||||
        if (m_context.fix_unbound_vars()) {
 | 
					        if (m_context.fix_unbound_vars()) {
 | 
				
			||||||
            m_rm.fix_unbound_vars(res, true);
 | 
					            m_rm.fix_unbound_vars(res, true);
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -40,24 +40,25 @@ namespace datalog {
 | 
				
			||||||
        bool           m_ready;
 | 
					        bool           m_ready;
 | 
				
			||||||
        unsigned       m_deltas[2];
 | 
					        unsigned       m_deltas[2];
 | 
				
			||||||
    public:
 | 
					    public:
 | 
				
			||||||
        rule_unifier(rule_manager& rm, context& ctx, ast_manager& m)
 | 
					        rule_unifier(context& ctx)
 | 
				
			||||||
            : m(m), m_rm(rm), m_context(ctx), m_interp_simplifier(ctx), m_subst(m), m_unif(m), m_ready(false) {}
 | 
					            : m(ctx.get_manager()), m_rm(ctx.get_rule_manager()), m_context(ctx), 
 | 
				
			||||||
 | 
					              m_interp_simplifier(ctx), m_subst(m), m_unif(m), m_ready(false) {}
 | 
				
			||||||
            
 | 
					            
 | 
				
			||||||
        /** Reset subtitution and unify tail tgt_idx of the target rule and the head of the src rule */
 | 
					        /** Reset subtitution and unify tail tgt_idx of the target rule and the head of the src rule */
 | 
				
			||||||
        bool unify_rules(const rule& tgt, unsigned tgt_idx, const rule& src);
 | 
					        bool unify_rules(rule const& tgt, unsigned tgt_idx, rule const& src);
 | 
				
			||||||
 | 
					
 | 
				
			||||||
        /**
 | 
					        /**
 | 
				
			||||||
           \brief Apply unifier to rules.
 | 
					           \brief Apply unifier to rules.
 | 
				
			||||||
           Return false if the resulting rule is a tautology (the interpreted tail is unsat).
 | 
					           Return false if the resulting rule is a tautology (the interpreted tail is unsat).
 | 
				
			||||||
        */
 | 
					        */
 | 
				
			||||||
        bool apply(rule& tgt, unsigned tgt_idx, rule& src, rule_ref& result);
 | 
					        bool apply(rule const& tgt, unsigned tgt_idx, rule const& src, rule_ref& result);
 | 
				
			||||||
 | 
					
 | 
				
			||||||
        void display(std::ostream& stm) { m_subst.display(stm, 2, m_deltas); }
 | 
					        void display(std::ostream& stm) { m_subst.display(stm, 2, m_deltas); }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
        /**
 | 
					        /**
 | 
				
			||||||
           Retrieve substitutions for src/tgt. (second argument of unify_rules).
 | 
					           Retrieve substitutions for src/tgt. (second argument of unify_rules).
 | 
				
			||||||
        */
 | 
					        */
 | 
				
			||||||
        expr_ref_vector get_rule_subst(const rule& r, bool is_tgt);
 | 
					        expr_ref_vector get_rule_subst(rule const& r, bool is_tgt);
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    private:
 | 
					    private:
 | 
				
			||||||
        void apply(app * a, bool is_tgt, app_ref& res);
 | 
					        void apply(app * a, bool is_tgt, app_ref& res);
 | 
				
			||||||
| 
						 | 
					@ -66,7 +67,7 @@ namespace datalog {
 | 
				
			||||||
           Apply substitution to a rule tail. Tail with skipped_index is skipped, 
 | 
					           Apply substitution to a rule tail. Tail with skipped_index is skipped, 
 | 
				
			||||||
           unless skipped_index is equal to UINT_MAX
 | 
					           unless skipped_index is equal to UINT_MAX
 | 
				
			||||||
        */        
 | 
					        */        
 | 
				
			||||||
        void apply(rule& r, bool is_tgt, unsigned skipped_index, app_ref_vector& res, 
 | 
					        void apply(rule const& r, bool is_tgt, unsigned skipped_index, app_ref_vector& res, 
 | 
				
			||||||
                   svector<bool>& res_neg);
 | 
					                   svector<bool>& res_neg);
 | 
				
			||||||
        
 | 
					        
 | 
				
			||||||
    };
 | 
					    };
 | 
				
			||||||
| 
						 | 
					@ -178,7 +179,7 @@ namespace datalog {
 | 
				
			||||||
            m_simp(m_context.get_rewriter()),
 | 
					            m_simp(m_context.get_rewriter()),
 | 
				
			||||||
            m_pinned(m_rm),
 | 
					            m_pinned(m_rm),
 | 
				
			||||||
            m_inlined_rules(m_context),
 | 
					            m_inlined_rules(m_context),
 | 
				
			||||||
            m_unifier(ctx.get_rule_manager(), ctx, m),
 | 
					            m_unifier(ctx),
 | 
				
			||||||
            m_mc(0),
 | 
					            m_mc(0),
 | 
				
			||||||
            m_pc(0),
 | 
					            m_pc(0),
 | 
				
			||||||
            m_head_index(m),
 | 
					            m_head_index(m),
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -259,7 +259,7 @@ namespace datalog {
 | 
				
			||||||
            rm(ctx.get_rule_manager()),
 | 
					            rm(ctx.get_rule_manager()),
 | 
				
			||||||
            m_pinned_rules(rm),
 | 
					            m_pinned_rules(rm),
 | 
				
			||||||
            m_pinned_exprs(m),
 | 
					            m_pinned_exprs(m),
 | 
				
			||||||
            m_unifier(rm, ctx, m) {}
 | 
					            m_unifier(ctx) {}
 | 
				
			||||||
 | 
					
 | 
				
			||||||
        void insert(rule* orig_rule, rule* slice_rule, unsigned sz, unsigned const* renaming) {
 | 
					        void insert(rule* orig_rule, rule* slice_rule, unsigned sz, unsigned const* renaming) {
 | 
				
			||||||
            m_rule2slice.insert(orig_rule, slice_rule);
 | 
					            m_rule2slice.insert(orig_rule, slice_rule);
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
							
								
								
									
										72
									
								
								lib/dl_mk_unfold.cpp
									
										
									
									
									
										Normal file
									
								
							
							
						
						
									
										72
									
								
								lib/dl_mk_unfold.cpp
									
										
									
									
									
										Normal file
									
								
							| 
						 | 
					@ -0,0 +1,72 @@
 | 
				
			||||||
 | 
					/*++
 | 
				
			||||||
 | 
					Copyright (c) 2012 Microsoft Corporation
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Module Name:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    dl_mk_unfold.cpp
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Abstract:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    Unfold rules once, return the unfolded set of rules.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Author:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    Nikolaj Bjorner (nbjorner) 2012-10-15
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Revision History:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					--*/
 | 
				
			||||||
 | 
					#include "dl_mk_unfold.h"
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					namespace datalog {
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    mk_unfold::mk_unfold(context& ctx):
 | 
				
			||||||
 | 
					        rule_transformer::plugin(0, false),
 | 
				
			||||||
 | 
					        m_ctx(ctx),
 | 
				
			||||||
 | 
					        m(ctx.get_manager()),
 | 
				
			||||||
 | 
					        rm(ctx.get_rule_manager()),
 | 
				
			||||||
 | 
					        m_unify(ctx)
 | 
				
			||||||
 | 
					    {}
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    void mk_unfold::expand_tail(rule& r, unsigned tail_idx, rule_set const& src, rule_set& dst) {
 | 
				
			||||||
 | 
					        SASSERT(tail_idx <= r.get_uninterpreted_tail_size());
 | 
				
			||||||
 | 
					        if (tail_idx == r.get_uninterpreted_tail_size()) {
 | 
				
			||||||
 | 
					            dst.add_rule(&r);
 | 
				
			||||||
 | 
					        }
 | 
				
			||||||
 | 
					        else {
 | 
				
			||||||
 | 
					            func_decl* p = r.get_decl(tail_idx);
 | 
				
			||||||
 | 
					            rule_vector const& p_rules = src.get_predicate_rules(p);
 | 
				
			||||||
 | 
					            rule_ref new_rule(rm);
 | 
				
			||||||
 | 
					            for (unsigned i = 0; i < p_rules.size(); ++i) {
 | 
				
			||||||
 | 
					                rule const& r2 = *p_rules[i];
 | 
				
			||||||
 | 
					                if (m_unify.unify_rules(r, tail_idx, r2) &&
 | 
				
			||||||
 | 
					                    m_unify.apply(r, tail_idx, r2, new_rule)) {
 | 
				
			||||||
 | 
					                    expr_ref_vector s1 = m_unify.get_rule_subst(r, true);
 | 
				
			||||||
 | 
					                    expr_ref_vector s2 = m_unify.get_rule_subst(r2, false);
 | 
				
			||||||
 | 
					                    resolve_rule(m_pc, r, r2, tail_idx, s1, s2, *new_rule.get());
 | 
				
			||||||
 | 
					                    expand_tail(*new_rule.get(), tail_idx+r2.get_uninterpreted_tail_size(), src, dst);
 | 
				
			||||||
 | 
					                }
 | 
				
			||||||
 | 
					            }
 | 
				
			||||||
 | 
					        }
 | 
				
			||||||
 | 
					    }
 | 
				
			||||||
 | 
					        
 | 
				
			||||||
 | 
					    rule_set * mk_unfold::operator()(rule_set const & source, model_converter_ref& mc, proof_converter_ref& pc) {
 | 
				
			||||||
 | 
					        m_pc = 0;
 | 
				
			||||||
 | 
					        ref<replace_proof_converter> rpc;
 | 
				
			||||||
 | 
					        if (pc) {
 | 
				
			||||||
 | 
					            rpc = alloc(replace_proof_converter, m);
 | 
				
			||||||
 | 
					            m_pc = rpc.get();
 | 
				
			||||||
 | 
					        }
 | 
				
			||||||
 | 
					        rule_set* rules = alloc(rule_set, m_ctx);
 | 
				
			||||||
 | 
					        rule_set::iterator it = source.begin(), end = source.end();
 | 
				
			||||||
 | 
					        for (; it != end; ++it) {
 | 
				
			||||||
 | 
					            expand_tail(**it, 0, source, *rules);
 | 
				
			||||||
 | 
					        }
 | 
				
			||||||
 | 
					        if (pc) {
 | 
				
			||||||
 | 
					            pc = concat(pc.get(), rpc.get());
 | 
				
			||||||
 | 
					        }
 | 
				
			||||||
 | 
					        return rules;
 | 
				
			||||||
 | 
					    }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					};
 | 
				
			||||||
 | 
					
 | 
				
			||||||
							
								
								
									
										54
									
								
								lib/dl_mk_unfold.h
									
										
									
									
									
										Normal file
									
								
							
							
						
						
									
										54
									
								
								lib/dl_mk_unfold.h
									
										
									
									
									
										Normal file
									
								
							| 
						 | 
					@ -0,0 +1,54 @@
 | 
				
			||||||
 | 
					/*++
 | 
				
			||||||
 | 
					Copyright (c) 2012 Microsoft Corporation
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Module Name:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    dl_mk_unfold.h
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Abstract:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    Unfold rules once, return the unfolded set of rules.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Author:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    Nikolaj Bjorner (nbjorner) 2012-10-15
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Revision History:
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					--*/
 | 
				
			||||||
 | 
					#ifndef _DL_MK_UNFOLD_H_
 | 
				
			||||||
 | 
					#define _DL_MK_UNFOLD_H_
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					#include"dl_context.h"
 | 
				
			||||||
 | 
					#include"dl_rule_set.h"
 | 
				
			||||||
 | 
					#include"uint_set.h"
 | 
				
			||||||
 | 
					#include"dl_rule_transformer.h"
 | 
				
			||||||
 | 
					#include"dl_mk_rule_inliner.h"
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					namespace datalog {
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    /**
 | 
				
			||||||
 | 
					       \brief Implements an unfolding transformation.
 | 
				
			||||||
 | 
					    */
 | 
				
			||||||
 | 
					    class mk_unfold : public rule_transformer::plugin {
 | 
				
			||||||
 | 
					        context&        m_ctx;
 | 
				
			||||||
 | 
					        ast_manager&    m;
 | 
				
			||||||
 | 
					        rule_manager&   rm;
 | 
				
			||||||
 | 
					        rule_unifier    m_unify;
 | 
				
			||||||
 | 
					        replace_proof_converter* m_pc;
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					        void expand_tail(rule& r, unsigned tail_idx, rule_set const& src, rule_set& dst);
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    public:
 | 
				
			||||||
 | 
					        /**
 | 
				
			||||||
 | 
					           \brief Create unfold rule transformer.
 | 
				
			||||||
 | 
					         */
 | 
				
			||||||
 | 
					        mk_unfold(context & ctx);
 | 
				
			||||||
 | 
					        
 | 
				
			||||||
 | 
					        rule_set * operator()(rule_set const & source, model_converter_ref& mc, proof_converter_ref& pc);
 | 
				
			||||||
 | 
					    };
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					};
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					#endif /* _DL_MK_UNFOLD_H_ */
 | 
				
			||||||
 | 
					
 | 
				
			||||||
| 
						 | 
					@ -615,8 +615,8 @@ namespace datalog {
 | 
				
			||||||
        }
 | 
					        }
 | 
				
			||||||
    }
 | 
					    }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    void resolve_rule(replace_proof_converter* pc, rule& r1, rule& r2, unsigned idx, 
 | 
					    void resolve_rule(replace_proof_converter* pc, rule const& r1, rule const& r2, unsigned idx, 
 | 
				
			||||||
                      expr_ref_vector const& s1, expr_ref_vector const& s2, rule& res) {
 | 
					                      expr_ref_vector const& s1, expr_ref_vector const& s2, rule const& res) {
 | 
				
			||||||
        if (!pc) return;
 | 
					        if (!pc) return;
 | 
				
			||||||
        ast_manager& m = s1.get_manager();
 | 
					        ast_manager& m = s1.get_manager();
 | 
				
			||||||
        dl_decl_util util(m);
 | 
					        dl_decl_util util(m);
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -467,8 +467,8 @@ namespace datalog {
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    void del_rule(horn_subsume_model_converter* mc, rule& r);
 | 
					    void del_rule(horn_subsume_model_converter* mc, rule& r);
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    void resolve_rule(replace_proof_converter* pc, rule& r1, rule& r2, unsigned idx, 
 | 
					    void resolve_rule(replace_proof_converter* pc, rule const& r1, rule const& r2, unsigned idx, 
 | 
				
			||||||
                      expr_ref_vector const& s1, expr_ref_vector const& s2, rule& res);
 | 
					                      expr_ref_vector const& s1, expr_ref_vector const& s2, rule const& res);
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    model_converter* mk_skip_model_converter();
 | 
					    model_converter* mk_skip_model_converter();
 | 
				
			||||||
 | 
					
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -559,6 +559,7 @@
 | 
				
			||||||
    <ClCompile Include="dl_mk_simple_joins.cpp" />
 | 
					    <ClCompile Include="dl_mk_simple_joins.cpp" />
 | 
				
			||||||
    <ClCompile Include="dl_mk_subsumption_checker.cpp" />
 | 
					    <ClCompile Include="dl_mk_subsumption_checker.cpp" />
 | 
				
			||||||
    <ClCompile Include="dl_mk_unbound_compressor.cpp" />
 | 
					    <ClCompile Include="dl_mk_unbound_compressor.cpp" />
 | 
				
			||||||
 | 
					    <ClCompile Include="dl_mk_unfold.cpp" />
 | 
				
			||||||
    <ClCompile Include="dl_relation_manager.cpp" />
 | 
					    <ClCompile Include="dl_relation_manager.cpp" />
 | 
				
			||||||
    <ClCompile Include="dl_product_relation.cpp" />
 | 
					    <ClCompile Include="dl_product_relation.cpp" />
 | 
				
			||||||
    <ClCompile Include="dl_rule.cpp" />
 | 
					    <ClCompile Include="dl_rule.cpp" />
 | 
				
			||||||
| 
						 | 
					@ -1009,6 +1010,7 @@
 | 
				
			||||||
    <ClInclude Include="dl_mk_similarity_compressor.h" />
 | 
					    <ClInclude Include="dl_mk_similarity_compressor.h" />
 | 
				
			||||||
    <ClInclude Include="dl_mk_subsumption_checker.h" />
 | 
					    <ClInclude Include="dl_mk_subsumption_checker.h" />
 | 
				
			||||||
    <ClInclude Include="dl_mk_unbound_compressor.h" />
 | 
					    <ClInclude Include="dl_mk_unbound_compressor.h" />
 | 
				
			||||||
 | 
					    <ClInclude Include="dl_mk_unfold.h" />
 | 
				
			||||||
    <ClInclude Include="dl_relation_manager.h" />
 | 
					    <ClInclude Include="dl_relation_manager.h" />
 | 
				
			||||||
    <ClInclude Include="dl_product_relation.h" />
 | 
					    <ClInclude Include="dl_product_relation.h" />
 | 
				
			||||||
    <ClInclude Include="dl_rule_subsumption_index.h" />
 | 
					    <ClInclude Include="dl_rule_subsumption_index.h" />
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -1026,7 +1026,7 @@ namespace pdr {
 | 
				
			||||||
        ast_manager& m = pt.get_manager();
 | 
					        ast_manager& m = pt.get_manager();
 | 
				
			||||||
        datalog::context& dctx = ctx.get_context();
 | 
					        datalog::context& dctx = ctx.get_context();
 | 
				
			||||||
        datalog::rule_manager& rm = dctx.get_rule_manager();
 | 
					        datalog::rule_manager& rm = dctx.get_rule_manager();
 | 
				
			||||||
        datalog::rule_unifier unifier(rm, dctx, m);
 | 
					        datalog::rule_unifier unifier(dctx);
 | 
				
			||||||
        datalog::dl_decl_util util(m);
 | 
					        datalog::dl_decl_util util(m);
 | 
				
			||||||
        datalog::rule_ref r0(rm), r1(rm);
 | 
					        datalog::rule_ref r0(rm), r1(rm);
 | 
				
			||||||
        obj_map<expr, proof*> cache;
 | 
					        obj_map<expr, proof*> cache;
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -30,6 +30,7 @@ Revision History:
 | 
				
			||||||
#include "pdr_dl_interface.h"
 | 
					#include "pdr_dl_interface.h"
 | 
				
			||||||
#include "dl_rule_set.h"
 | 
					#include "dl_rule_set.h"
 | 
				
			||||||
#include "dl_mk_slice.h"
 | 
					#include "dl_mk_slice.h"
 | 
				
			||||||
 | 
					#include "dl_mk_unfold.h"
 | 
				
			||||||
 | 
					
 | 
				
			||||||
using namespace pdr;
 | 
					using namespace pdr;
 | 
				
			||||||
 | 
					
 | 
				
			||||||
| 
						 | 
					@ -117,6 +118,17 @@ lbool dl_interface::query(expr * query) {
 | 
				
			||||||
        m_ctx.set_output_predicate(query_pred);
 | 
					        m_ctx.set_output_predicate(query_pred);
 | 
				
			||||||
    }
 | 
					    }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    if (m_ctx.get_params().get_uint(":unfold-rules",0) > 0) {
 | 
				
			||||||
 | 
					        unsigned num_unfolds = m_ctx.get_params().get_uint(":unfold-rules",0);
 | 
				
			||||||
 | 
					        datalog::rule_transformer transformer(m_ctx);
 | 
				
			||||||
 | 
					        datalog::mk_unfold* unfold = alloc(datalog::mk_unfold, m_ctx);
 | 
				
			||||||
 | 
					        transformer.register_plugin(unfold);
 | 
				
			||||||
 | 
					        while (num_unfolds > 0) {
 | 
				
			||||||
 | 
					            m_ctx.transform_rules(transformer, mc, pc);        
 | 
				
			||||||
 | 
					            --num_unfolds;
 | 
				
			||||||
 | 
					        }
 | 
				
			||||||
 | 
					    }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    IF_VERBOSE(2, m_ctx.display_rules(verbose_stream()););
 | 
					    IF_VERBOSE(2, m_ctx.display_rules(verbose_stream()););
 | 
				
			||||||
    m_pdr_rules.add_rules(m_ctx.get_rules());
 | 
					    m_pdr_rules.add_rules(m_ctx.get_rules());
 | 
				
			||||||
    m_pdr_rules.close();
 | 
					    m_pdr_rules.close();
 | 
				
			||||||
| 
						 | 
					@ -187,6 +199,7 @@ void dl_interface::collect_params(param_descrs& p) {
 | 
				
			||||||
    p.insert(":generate-proof-trace", CPK_BOOL, "PDR: (default false) trace for 'sat' answer as proof object");
 | 
					    p.insert(":generate-proof-trace", CPK_BOOL, "PDR: (default false) trace for 'sat' answer as proof object");
 | 
				
			||||||
    p.insert(":inline-proofs", CPK_BOOL, "PDR: (default true) run PDR with proof mode turned on and extract Farkas coefficients directly (instead of creating a separate proof object when extracting coefficients)"); 
 | 
					    p.insert(":inline-proofs", CPK_BOOL, "PDR: (default true) run PDR with proof mode turned on and extract Farkas coefficients directly (instead of creating a separate proof object when extracting coefficients)"); 
 | 
				
			||||||
    p.insert(":flexible-trace", CPK_BOOL, "PDR: (default false) allow PDR generate long counter-examples by extending candidate trace within search area");
 | 
					    p.insert(":flexible-trace", CPK_BOOL, "PDR: (default false) allow PDR generate long counter-examples by extending candidate trace within search area");
 | 
				
			||||||
 | 
					    p.insert(":unfold-rules", CPK_UINT, "PDR: (default 0) unfold rules statically using iterative squarring");
 | 
				
			||||||
    PRIVATE_PARAMS(p.insert(":use-farkas-model", CPK_BOOL, "PDR: (default false) enable using Farkas generalization through model propagation"););
 | 
					    PRIVATE_PARAMS(p.insert(":use-farkas-model", CPK_BOOL, "PDR: (default false) enable using Farkas generalization through model propagation"););
 | 
				
			||||||
    PRIVATE_PARAMS(p.insert(":use-precondition-generalizer", CPK_BOOL, "PDR: (default false) enable generalizations from weakest pre-conditions"););
 | 
					    PRIVATE_PARAMS(p.insert(":use-precondition-generalizer", CPK_BOOL, "PDR: (default false) enable generalizations from weakest pre-conditions"););
 | 
				
			||||||
    PRIVATE_PARAMS(p.insert(":use-multicore-generalizer", CPK_BOOL, "PDR: (default false) extract multiple cores for blocking states"););
 | 
					    PRIVATE_PARAMS(p.insert(":use-multicore-generalizer", CPK_BOOL, "PDR: (default false) extract multiple cores for blocking states"););
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -683,123 +683,3 @@ namespace pdr {
 | 
				
			||||||
    }
 | 
					    }
 | 
				
			||||||
 | 
					
 | 
				
			||||||
};
 | 
					};
 | 
				
			||||||
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
#if 0
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    void model_farkas_generalizer::extract_eqs(expr_ref_vector const& lits, eqs& eqs) {
 | 
					 | 
				
			||||||
        expr* e;
 | 
					 | 
				
			||||||
        rational vl;
 | 
					 | 
				
			||||||
        bool is_int;
 | 
					 | 
				
			||||||
        for (unsigned i = 0; i < lits.size(); ++i) {
 | 
					 | 
				
			||||||
            if (is_eq(lits[i], e, vl, is_int)) {                
 | 
					 | 
				
			||||||
                eqs.push_back(std::make_pair(e, vl));
 | 
					 | 
				
			||||||
            }
 | 
					 | 
				
			||||||
        }        
 | 
					 | 
				
			||||||
    }
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    void model_farkas_generalizer::solve_eqs(expr_ref& A, expr_ref_vector& other, eqs const& o_eqs) {
 | 
					 | 
				
			||||||
        if (o_eqs.empty()) {
 | 
					 | 
				
			||||||
            return;
 | 
					 | 
				
			||||||
        }
 | 
					 | 
				
			||||||
        
 | 
					 | 
				
			||||||
        expr_substitution sub(m);        
 | 
					 | 
				
			||||||
        scoped_ptr<expr_replacer> rep = mk_default_expr_replacer(m);
 | 
					 | 
				
			||||||
        rep->set_substitution(&sub);
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
        expr* e1;
 | 
					 | 
				
			||||||
        rational vl1;
 | 
					 | 
				
			||||||
        bool is_int;
 | 
					 | 
				
			||||||
        for (unsigned i = 0; i < other.size(); ++i) {
 | 
					 | 
				
			||||||
            if (is_eq(other[i].get(), e1, vl1, is_int)) {
 | 
					 | 
				
			||||||
                unsigned k = m_random(o_eqs.size());
 | 
					 | 
				
			||||||
                for (unsigned j = 0; j < o_eqs.size(); ++j) {
 | 
					 | 
				
			||||||
                    unsigned l = (j+k)%o_eqs.size();
 | 
					 | 
				
			||||||
                    expr* e2 = o_eqs[l].first;
 | 
					 | 
				
			||||||
                    if (m.get_sort(e1) == m.get_sort(e2)) {
 | 
					 | 
				
			||||||
                        sub.insert(e1, mk_offset(e2, vl1-o_eqs[l].second));
 | 
					 | 
				
			||||||
                        other[i] = m.mk_true();
 | 
					 | 
				
			||||||
                        break;
 | 
					 | 
				
			||||||
                    }
 | 
					 | 
				
			||||||
                }                
 | 
					 | 
				
			||||||
            }
 | 
					 | 
				
			||||||
        }
 | 
					 | 
				
			||||||
        (*rep)(A);
 | 
					 | 
				
			||||||
        other.push_back(A);
 | 
					 | 
				
			||||||
        A = m.mk_and(other.size(), other.c_ptr());
 | 
					 | 
				
			||||||
    }
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    expr* model_farkas_generalizer::mk_offset(expr* e, rational const& r) {
 | 
					 | 
				
			||||||
        if (r.is_zero()) {
 | 
					 | 
				
			||||||
            return e;
 | 
					 | 
				
			||||||
        }
 | 
					 | 
				
			||||||
        else {
 | 
					 | 
				
			||||||
            return a.mk_add(e, a.mk_numeral(r, a.is_int(e)));
 | 
					 | 
				
			||||||
        }
 | 
					 | 
				
			||||||
    }
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    void model_farkas_generalizer::connect_vars(ptr_vector<expr>& vars, vector<rational>& vals, expr_ref_vector& lits) {
 | 
					 | 
				
			||||||
        switch(vars.size()) {
 | 
					 | 
				
			||||||
        case 0: 
 | 
					 | 
				
			||||||
            break;
 | 
					 | 
				
			||||||
        case 1: 
 | 
					 | 
				
			||||||
            lits.push_back(m.mk_eq(vars[0], a.mk_numeral(vals[0], a.is_int(vars[0])))); 
 | 
					 | 
				
			||||||
            break;
 | 
					 | 
				
			||||||
        case 2:
 | 
					 | 
				
			||||||
            lits.push_back(m.mk_eq(vars[0], a.mk_numeral(vals[0], a.is_int(vars[0])))); 
 | 
					 | 
				
			||||||
            lits.push_back(m.mk_eq(vars[0], mk_offset(vars[1], vals[0]-vals[1])));
 | 
					 | 
				
			||||||
            break;
 | 
					 | 
				
			||||||
        default: {
 | 
					 | 
				
			||||||
            ptr_vector<expr> new_vars;
 | 
					 | 
				
			||||||
            vector<rational> new_vals;
 | 
					 | 
				
			||||||
            unsigned j, i = m_random(vars.size());
 | 
					 | 
				
			||||||
            new_vars.push_back(vars[i]);
 | 
					 | 
				
			||||||
            new_vals.push_back(vals[i]);
 | 
					 | 
				
			||||||
            lits.push_back(m.mk_eq(vars[i], a.mk_numeral(vals[i], a.is_int(vars[i])))); 
 | 
					 | 
				
			||||||
            vars.erase(vars.begin() + i);
 | 
					 | 
				
			||||||
            vals.erase(vals.begin() + i);
 | 
					 | 
				
			||||||
            while (!vars.empty()) {
 | 
					 | 
				
			||||||
                i = m_random(vars.size());
 | 
					 | 
				
			||||||
                j = m_random(new_vars.size());
 | 
					 | 
				
			||||||
                lits.push_back(m.mk_eq(vars[i], mk_offset(new_vars[j], vals[i]-new_vals[j])));
 | 
					 | 
				
			||||||
                new_vars.push_back(vars[i]);
 | 
					 | 
				
			||||||
                new_vals.push_back(vals[i]);
 | 
					 | 
				
			||||||
                vars.erase(vars.begin() + i);
 | 
					 | 
				
			||||||
                vals.erase(vals.begin() + i);
 | 
					 | 
				
			||||||
            }
 | 
					 | 
				
			||||||
            break;
 | 
					 | 
				
			||||||
        }
 | 
					 | 
				
			||||||
        }
 | 
					 | 
				
			||||||
    }
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    bool model_farkas_generalizer::is_eq(expr* e, expr*& r, rational& vl, bool& is_int) {
 | 
					 | 
				
			||||||
        expr* r2 = 0;
 | 
					 | 
				
			||||||
        return 
 | 
					 | 
				
			||||||
            (m.is_eq(e, r, r2) && a.is_numeral(r2, vl, is_int)) ||
 | 
					 | 
				
			||||||
            (m.is_eq(e, r2, r) && a.is_numeral(r2, vl, is_int));
 | 
					 | 
				
			||||||
    }
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    void model_farkas_generalizer::relativize(expr_ref_vector& literals) {
 | 
					 | 
				
			||||||
        ast_manager& m = literals.get_manager();
 | 
					 | 
				
			||||||
        ptr_vector<expr> ints, reals;
 | 
					 | 
				
			||||||
        vector<rational> int_vals, real_vals;
 | 
					 | 
				
			||||||
        expr* e;
 | 
					 | 
				
			||||||
        rational vl;
 | 
					 | 
				
			||||||
        bool is_int;
 | 
					 | 
				
			||||||
        for (unsigned i = 0; i < literals.size(); ) {
 | 
					 | 
				
			||||||
            if (is_eq(literals[i].get(), e, vl, is_int)) {                
 | 
					 | 
				
			||||||
                (is_int?ints:reals).push_back(e); 
 | 
					 | 
				
			||||||
                (is_int?int_vals:real_vals).push_back(vl);
 | 
					 | 
				
			||||||
                literals[i] = literals.back();
 | 
					 | 
				
			||||||
                literals.pop_back();
 | 
					 | 
				
			||||||
            }
 | 
					 | 
				
			||||||
            else {
 | 
					 | 
				
			||||||
                ++i;
 | 
					 | 
				
			||||||
            }
 | 
					 | 
				
			||||||
        }
 | 
					 | 
				
			||||||
        connect_vars(ints, int_vals, literals);
 | 
					 | 
				
			||||||
        connect_vars(reals, real_vals, literals);
 | 
					 | 
				
			||||||
    }
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
#endif
 | 
					 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
| 
						 | 
					@ -110,7 +110,7 @@ namespace smt2 {
 | 
				
			||||||
        m_string.reset();
 | 
					        m_string.reset();
 | 
				
			||||||
        m_string.push_back(curr());
 | 
					        m_string.push_back(curr());
 | 
				
			||||||
        next();
 | 
					        next();
 | 
				
			||||||
        read_symbol_core();
 | 
					        return read_symbol_core();
 | 
				
			||||||
    }
 | 
					    }
 | 
				
			||||||
    
 | 
					    
 | 
				
			||||||
    scanner::token scanner::read_number() {
 | 
					    scanner::token scanner::read_number() {
 | 
				
			||||||
| 
						 | 
					
 | 
				
			||||||
		Loading…
	
	Add table
		Add a link
		
	
		Reference in a new issue