feat(library/blast/unit): preprocessor placeholder
This commit is contained in:
parent
0df4556eb5
commit
6e478696d2
19 changed files with 352 additions and 46 deletions
|
@ -7,4 +7,4 @@ prelude
|
|||
import init.datatypes init.reserved_notation init.tactic init.logic
|
||||
import init.relation init.wf init.nat init.wf_k init.prod
|
||||
import init.bool init.num init.sigma init.measurable init.setoid init.quot
|
||||
import init.funext init.function init.subtype init.classical
|
||||
import init.funext init.function init.subtype init.classical init.simplifier
|
||||
|
|
85
library/init/simplifier.lean
Normal file
85
library/init/simplifier.lean
Normal file
|
@ -0,0 +1,85 @@
|
|||
/-
|
||||
Copyright (c) 2015 Daniel Selsam. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Daniel Selsam
|
||||
-/
|
||||
prelude
|
||||
import init.logic
|
||||
|
||||
namespace simplifier
|
||||
|
||||
namespace unit_simp
|
||||
open eq.ops
|
||||
|
||||
-- TODO(dhs): prove these lemmas elsewhere and only gather the
|
||||
-- [simp] attributes here
|
||||
|
||||
variables {A B C : Prop}
|
||||
|
||||
lemma and_imp [simp] : (A ∧ B → C) ↔ (A → B → C) :=
|
||||
iff.intro (assume H a b, H (and.intro a b))
|
||||
(assume H ab, H (and.left ab) (and.right ab))
|
||||
|
||||
lemma or_imp [simp] : (A ∨ B → C) ↔ ((A → C) ∧ (B → C)) :=
|
||||
iff.intro (assume H, and.intro (assume a, H (or.inl a))
|
||||
(assume b, H (or.inr b)))
|
||||
(assume H ab, and.rec_on H
|
||||
(assume Hac Hbc, or.rec_on ab Hac Hbc))
|
||||
|
||||
lemma imp_and [simp] : (A → B ∧ C) ↔ ((A → B) ∧ (A → C)) :=
|
||||
iff.intro (assume H, and.intro (assume a, and.left (H a))
|
||||
(assume a, and.right (H a)))
|
||||
(assume H a, and.rec_on H
|
||||
(assume Hab Hac, and.intro (Hab a) (Hac a)))
|
||||
|
||||
-- TODO(dhs, leo): do we want to pre-process away the [iff]s?
|
||||
/-
|
||||
lemma iff_and_imp [simp] : ((A ↔ B) → C) ↔ (((A → B) ∧ (B → A)) → C) :=
|
||||
iff.intro (assume H1 H2, and.rec_on H2 (assume ab ba, H1 (iff.intro ab ba)))
|
||||
(assume H1 H2, H1 (and.intro (iff.elim_left H2) (iff.elim_right H2)))
|
||||
-/
|
||||
|
||||
lemma a_of_a [simp] : (A → A) ↔ true :=
|
||||
iff.intro (assume H, trivial)
|
||||
(assume t a, a)
|
||||
|
||||
lemma not_true_of_false [simp] : ¬ true ↔ false :=
|
||||
iff.intro (assume H, H trivial)
|
||||
(assume f, false.rec (¬ true) f)
|
||||
|
||||
lemma imp_true [simp] : (A → true) ↔ true :=
|
||||
iff.intro (assume H, trivial)
|
||||
(assume t a, trivial)
|
||||
|
||||
lemma true_imp [simp] : (true → A) ↔ A :=
|
||||
iff.intro (assume H, H trivial)
|
||||
(assume a t, a)
|
||||
|
||||
lemma fold_not [simp] : (A → false) ↔ ¬ A :=
|
||||
iff.intro id id
|
||||
|
||||
lemma false_imp [simp] : (false → A) ↔ true :=
|
||||
iff.intro (assume H, trivial)
|
||||
(assume t f, false.rec A f)
|
||||
|
||||
lemma ite_and [simp] [A_dec : decidable A] : ite A B C ↔ ((A → B) ∧ (¬ A → C)) :=
|
||||
iff.intro (assume H, and.intro (assume a, implies_of_if_pos H a)
|
||||
(assume a, implies_of_if_neg H a))
|
||||
(assume H, and.rec_on H
|
||||
(assume Hab Hnac, decidable.rec_on A_dec
|
||||
(assume a,
|
||||
assert rw : @decidable.inl A a = A_dec, from
|
||||
subsingleton.rec_on (subsingleton_decidable A)
|
||||
(assume H, H (@decidable.inl A a) A_dec),
|
||||
by rewrite [rw, if_pos a] ; exact Hab a)
|
||||
(assume na,
|
||||
assert rw : @decidable.inr A na = A_dec, from
|
||||
subsingleton.rec_on (subsingleton_decidable A)
|
||||
(assume H, H (@decidable.inr A na) A_dec),
|
||||
by rewrite [rw, if_neg na] ; exact Hnac na)))
|
||||
|
||||
end unit_simp
|
||||
|
||||
|
||||
|
||||
end simplifier
|
|
@ -15,7 +15,7 @@ Author: Leonardo de Moura
|
|||
#include "library/blast/backward/backward_strategy.h"
|
||||
#include "library/blast/forward/forward_actions.h"
|
||||
#include "library/blast/forward/ematch.h"
|
||||
#include "library/blast/unit/unit_action.h"
|
||||
#include "library/blast/unit/unit_actions.h"
|
||||
#include "library/blast/no_confusion_action.h"
|
||||
#include "library/blast/simplifier/simplifier_actions.h"
|
||||
#include "library/blast/recursor_action.h"
|
||||
|
@ -31,6 +31,7 @@ class simple_strategy : public strategy {
|
|||
action_result hypothesis_pre_activation(hypothesis_idx hidx) override {
|
||||
Try(assumption_contradiction_actions(hidx));
|
||||
Try(simplify_hypothesis_action(hidx));
|
||||
Try(unit_preprocess(hidx));
|
||||
Try(no_confusion_action(hidx));
|
||||
TrySolve(assert_cc_action(hidx));
|
||||
Try(discard_action(hidx));
|
||||
|
@ -39,7 +40,7 @@ class simple_strategy : public strategy {
|
|||
}
|
||||
|
||||
action_result hypothesis_post_activation(hypothesis_idx hidx) override {
|
||||
Try(unit_action(hidx));
|
||||
Try(unit_propagate(hidx));
|
||||
Try(recursor_preprocess_action(hidx));
|
||||
return action_result::new_branch();
|
||||
}
|
||||
|
|
|
@ -134,6 +134,7 @@ static bool is_neg_app(expr const & e) {
|
|||
class simplifier {
|
||||
blast_tmp_type_context m_tmp_tctx;
|
||||
name m_rel;
|
||||
expr_predicate m_simp_pred;
|
||||
|
||||
simp_rule_sets m_srss;
|
||||
simp_rule_sets m_ctx_srss;
|
||||
|
@ -255,7 +256,7 @@ class simplifier {
|
|||
|
||||
|
||||
public:
|
||||
simplifier(name const & rel): m_rel(rel) { }
|
||||
simplifier(name const & rel, expr_predicate const & simp_pred): m_rel(rel), m_simp_pred(simp_pred) { }
|
||||
result operator()(expr const & e, simp_rule_sets const & srss) { return simplify(e, srss); }
|
||||
};
|
||||
|
||||
|
@ -383,6 +384,8 @@ result simplifier::simplify(expr const & e, bool is_root) {
|
|||
if (auto it = cache_lookup(e)) return *it;
|
||||
}
|
||||
|
||||
if (!m_simp_pred(e)) return result(e);
|
||||
|
||||
if (m_numerals && using_eq()) {
|
||||
if (auto r = simplify_numeral(e)) {
|
||||
return *r;
|
||||
|
@ -719,7 +722,8 @@ result simplifier::try_congr(expr const & e, congr_rule const & cr) {
|
|||
lean_verify(is_simp_relation(env(), m_type, h_rel, h_lhs, h_rhs) && is_constant(h_rel));
|
||||
{
|
||||
flet<name> set_name(m_rel, const_name(h_rel));
|
||||
flet<simp_rule_sets> set_ctx_srss(m_ctx_srss, add_to_srss(m_ctx_srss, ls));
|
||||
|
||||
flet<simp_rule_sets> set_ctx_srss(m_ctx_srss, m_contextual ? add_to_srss(m_ctx_srss, ls) : m_ctx_srss);
|
||||
|
||||
h_lhs = tmp_tctx->instantiate_uvars_mvars(h_lhs);
|
||||
lean_assert(!has_metavar(h_lhs));
|
||||
|
@ -756,9 +760,11 @@ bool simplifier::instantiate_emetas(blast_tmp_type_context & tmp_tctx, unsigned
|
|||
expr m_type = tmp_tctx->instantiate_uvars_mvars(tmp_tctx->infer(m));
|
||||
lean_assert(!has_metavar(m_type));
|
||||
|
||||
if (tmp_tctx->is_mvar_assigned(i)) return;
|
||||
|
||||
if (is_instance) {
|
||||
if (auto v = tmp_tctx->mk_class_instance(m_type)) {
|
||||
if (!tmp_tctx->force_assign(m, *v)) {
|
||||
if (!tmp_tctx->assign(m, *v)) {
|
||||
if (m_trace) {
|
||||
ios().get_diagnostic_channel() << "unable to assign instance for: " << m_type << "\n";
|
||||
}
|
||||
|
@ -1098,9 +1104,15 @@ void finalize_simplifier() {
|
|||
delete g_simplify_empty_namespace;
|
||||
}
|
||||
|
||||
/* Entry point */
|
||||
/* Entry points */
|
||||
static bool simplify_all_pred(expr const &) { return true; }
|
||||
|
||||
result simplify(name const & rel, expr const & e, simp_rule_sets const & srss) {
|
||||
return simplifier(rel)(e, srss);
|
||||
return simplifier(rel, simplify_all_pred)(e, srss);
|
||||
}
|
||||
|
||||
result simplify(name const & rel, expr const & e, simp_rule_sets const & srss, expr_predicate const & simp_pred) {
|
||||
return simplifier(rel, simp_pred)(e, srss);
|
||||
}
|
||||
|
||||
}}
|
||||
|
|
|
@ -37,7 +37,11 @@ public:
|
|||
};
|
||||
}
|
||||
|
||||
// TODO(dhs): put this outside of blast module
|
||||
typedef std::function<bool(expr const &)> expr_predicate; // NOLINT
|
||||
|
||||
simp::result simplify(name const & rel, expr const & e, simp_rule_sets const & srss);
|
||||
simp::result simplify(name const & rel, expr const & e, simp_rule_sets const & srss, expr_predicate const & simp_pred);
|
||||
|
||||
void initialize_simplifier();
|
||||
void finalize_simplifier();
|
||||
|
|
|
@ -1 +1 @@
|
|||
add_library(unit OBJECT init_module.cpp unit_action.cpp)
|
||||
add_library(unit OBJECT init_module.cpp unit_propagate.cpp unit_preprocess.cpp)
|
||||
|
|
|
@ -3,16 +3,18 @@ Copyright (c) 2015 Daniel Selsam. All rights reserved.
|
|||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Daniel Selsam
|
||||
*/
|
||||
#include "library/blast/unit/unit_action.h"
|
||||
#include "library/blast/unit/unit_actions.h"
|
||||
|
||||
namespace lean {
|
||||
namespace blast {
|
||||
|
||||
void initialize_unit_module() {
|
||||
initialize_unit_action();
|
||||
initialize_unit_propagate();
|
||||
initialize_unit_preprocess();
|
||||
}
|
||||
|
||||
void finalize_unit_module() {
|
||||
finalize_unit_action();
|
||||
finalize_unit_preprocess();
|
||||
finalize_unit_propagate();
|
||||
}
|
||||
}}
|
||||
|
|
|
@ -1,25 +0,0 @@
|
|||
/*
|
||||
Copyright (c) 2015 Daniel Selsam. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Daniel Selsam
|
||||
*/
|
||||
#pragma once
|
||||
#include "library/blast/action_result.h"
|
||||
|
||||
namespace lean {
|
||||
namespace blast {
|
||||
/* \brief The unit module handles lemmas of the form
|
||||
<tt>A_1 -> ... -> A_n -> B_1 \/ (B2 \/ ... \/ B_m)...)</tt>
|
||||
|
||||
Whenever all but one of the literals is present as a hypothesis with
|
||||
the appropriate polarity, we instantiate and resolve as necessary
|
||||
to conclude a new literal.
|
||||
|
||||
Remark: we assume that a pre-processing step will put lemmas
|
||||
into the above form when possible.
|
||||
*/
|
||||
action_result unit_action(unsigned hidx);
|
||||
|
||||
void initialize_unit_action();
|
||||
void finalize_unit_action();
|
||||
}}
|
42
src/library/blast/unit/unit_actions.h
Normal file
42
src/library/blast/unit/unit_actions.h
Normal file
|
@ -0,0 +1,42 @@
|
|||
/*
|
||||
Copyright (c) 2015 Daniel Selsam. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Daniel Selsam
|
||||
*/
|
||||
#pragma once
|
||||
#include "library/blast/action_result.h"
|
||||
|
||||
namespace lean {
|
||||
namespace blast {
|
||||
/* \brief The unit module handles lemmas of the form
|
||||
<tt>A_1 -> ... -> A_n -> B_1 \/ (B2 \/ ... \/ B_m)...)</tt>
|
||||
|
||||
Whenever all but one of the literals is present as a hypothesis with
|
||||
the appropriate polarity, we instantiate and resolve as necessary
|
||||
to conclude a new literal.
|
||||
|
||||
Remark: we assume that a pre-processing step will put lemmas
|
||||
into the above form when possible.
|
||||
*/
|
||||
action_result unit_propagate(unsigned hidx);
|
||||
|
||||
/* \brief We pre-process propositional lemmas to put them in the form
|
||||
expected by [unit_propagate]. In particular, we apply the following
|
||||
six transformations:
|
||||
|
||||
1. A ∧ B → C ==> { A → B → C }
|
||||
2. A ∨ B → C ==> { A → C, B → C }
|
||||
3. A → B ∧ C ==> { A → C, A → C }
|
||||
4. A ↔ B ==> { (A → B) ∧ (B → A) }
|
||||
5. Push negations inside ∧ and ∨ (when using classical)
|
||||
6. ite A B C ==> { A → B, ¬ A → C }
|
||||
*/
|
||||
action_result unit_preprocess(unsigned hidx);
|
||||
|
||||
void initialize_unit_propagate();
|
||||
void finalize_unit_propagate();
|
||||
|
||||
void initialize_unit_preprocess();
|
||||
void finalize_unit_preprocess();
|
||||
|
||||
}}
|
80
src/library/blast/unit/unit_preprocess.cpp
Normal file
80
src/library/blast/unit/unit_preprocess.cpp
Normal file
|
@ -0,0 +1,80 @@
|
|||
/*
|
||||
Copyright (c) 2015 Daniel Selsam. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Author: Daniel Selsam
|
||||
*/
|
||||
#include "kernel/instantiate.h"
|
||||
#include "kernel/abstract.h"
|
||||
#include "kernel/inductive/inductive.h"
|
||||
#include "library/constants.h"
|
||||
#include "library/util.h"
|
||||
#include "library/blast/blast.h"
|
||||
#include "library/blast/action_result.h"
|
||||
#include "library/blast/unit/unit_actions.h"
|
||||
#include "library/blast/proof_expr.h"
|
||||
#include "library/blast/choice_point.h"
|
||||
#include "library/blast/hypothesis.h"
|
||||
#include "library/blast/util.h"
|
||||
#include "library/blast/simplifier/simp_rule_set.h"
|
||||
#include "library/blast/simplifier/simplifier.h"
|
||||
#include "library/expr_lt.h"
|
||||
#include "util/rb_multi_map.h"
|
||||
|
||||
namespace lean {
|
||||
namespace blast {
|
||||
|
||||
static name * g_simplify_unit_simp_namespace = nullptr;
|
||||
static name * g_simplify_contextual = nullptr;
|
||||
|
||||
static bool is_propositional(expr const & e) {
|
||||
// TODO(dhs): This predicate will need to evolve, and will eventually be thrown out
|
||||
// when we re-implement this module without relying on the simplifier.
|
||||
if (!is_prop(e)) return false;
|
||||
if (is_pi(e) && closed(binding_body(e))) return true;
|
||||
if (is_and(e) || is_or(e) || is_not(env(), e) || is_ite(e) || is_iff(e)) return true;
|
||||
return false;
|
||||
}
|
||||
|
||||
action_result unit_preprocess(unsigned hidx) {
|
||||
// TODO(dhs, leo): do not rely on the simplifier, and provide a "safe mode" where we
|
||||
// introduce auxiliary hypotheses
|
||||
state & s = curr_state();
|
||||
if (s.has_target_forward_deps(hidx)) {
|
||||
// We currently do not try to simplify a hypothesis if other
|
||||
// hypotheses or target depends on it.
|
||||
return action_result::failed();
|
||||
}
|
||||
hypothesis const & h = s.get_hypothesis_decl(hidx);
|
||||
if (!is_prop(h.get_type())) {
|
||||
// We currently only simplify propositions.
|
||||
return action_result::failed();
|
||||
}
|
||||
|
||||
simp_rule_sets srss = get_simp_rule_sets(env(), ios(), *g_simplify_unit_simp_namespace);
|
||||
// TODO(dhs): disable contextual rewriting
|
||||
auto r = simplify(get_iff_name(), h.get_type(), srss, is_propositional);
|
||||
|
||||
if (r.get_new() == h.get_type()) return action_result::failed(); // did nothing
|
||||
expr new_h_proof;
|
||||
if (r.has_proof()) {
|
||||
new_h_proof = get_app_builder().mk_app(get_iff_mp_name(), r.get_proof(), h.get_self());
|
||||
} else {
|
||||
// they are definitionally equal
|
||||
new_h_proof = h.get_self();
|
||||
}
|
||||
s.mk_hypothesis(r.get_new(), new_h_proof);
|
||||
s.del_hypothesis(hidx);
|
||||
return action_result::new_branch();
|
||||
}
|
||||
|
||||
void initialize_unit_preprocess() {
|
||||
g_simplify_unit_simp_namespace = new name{"simplifier", "unit_simp"};
|
||||
g_simplify_contextual = new name{"simplify", "contextual"};
|
||||
}
|
||||
|
||||
void finalize_unit_preprocess() {
|
||||
delete g_simplify_contextual;
|
||||
delete g_simplify_unit_simp_namespace;
|
||||
}
|
||||
|
||||
}}
|
|
@ -10,7 +10,7 @@ Author: Daniel Selsam
|
|||
#include "library/util.h"
|
||||
#include "library/blast/blast.h"
|
||||
#include "library/blast/action_result.h"
|
||||
#include "library/blast/unit/unit_action.h"
|
||||
#include "library/blast/unit/unit_actions.h"
|
||||
#include "library/blast/proof_expr.h"
|
||||
#include "library/blast/choice_point.h"
|
||||
#include "library/blast/hypothesis.h"
|
||||
|
@ -106,11 +106,11 @@ public:
|
|||
}
|
||||
};
|
||||
|
||||
void initialize_unit_action() {
|
||||
void initialize_unit_propagate() {
|
||||
g_ext_id = register_branch_extension(new unit_branch_extension());
|
||||
}
|
||||
|
||||
void finalize_unit_action() { }
|
||||
void finalize_unit_propagate() { }
|
||||
|
||||
static unit_branch_extension & get_extension() {
|
||||
return static_cast<unit_branch_extension&>(curr_state().get_extension(g_ext_id));
|
||||
|
@ -285,7 +285,7 @@ action_result unit_fact(expr const & type) {
|
|||
else return action_result::failed();
|
||||
}
|
||||
|
||||
action_result unit_action(unsigned hidx) {
|
||||
action_result unit_propagate(unsigned hidx) {
|
||||
hypothesis const & h = curr_state().get_hypothesis_decl(hidx);
|
||||
expr type = whnf(h.get_type());
|
||||
if (is_lemma(type)) return unit_lemma(hidx, type, h.get_self());
|
|
@ -340,6 +340,20 @@ bool is_not(environment const & env, expr const & e, expr & a) {
|
|||
}
|
||||
}
|
||||
|
||||
bool is_not(environment const & env, expr const & e) {
|
||||
if (is_app(e)) {
|
||||
expr const & f = app_fn(e);
|
||||
if (!is_constant(f) || const_name(f) != get_not_name())
|
||||
return false;
|
||||
return true;
|
||||
} else if (is_pi(e)) {
|
||||
if (!is_false(env, binding_body(e))) return false;
|
||||
return true;
|
||||
} else {
|
||||
return false;
|
||||
}
|
||||
}
|
||||
|
||||
expr mk_not(type_checker & tc, expr const & e) {
|
||||
if (is_standard(tc.env())) {
|
||||
return mk_app(mk_constant(get_not_name()), e);
|
||||
|
@ -422,6 +436,16 @@ bool is_and(expr const & e, expr & arg1, expr & arg2) {
|
|||
}
|
||||
}
|
||||
|
||||
bool is_and(expr const & e) {
|
||||
if (get_app_fn(e) == *g_and) {
|
||||
buffer<expr> args; get_app_args(e, args);
|
||||
if (args.size() == 2) return true;
|
||||
else return false;
|
||||
} else {
|
||||
return false;
|
||||
}
|
||||
}
|
||||
|
||||
expr mk_and(expr const & a, expr const & b) {
|
||||
return mk_app(*g_and, a, b);
|
||||
}
|
||||
|
@ -494,13 +518,25 @@ bool is_ite(expr const & e, expr & c, expr & H, expr & A, expr & t, expr & f) {
|
|||
c = args[0]; H = args[1]; A = args[2]; t = args[3]; f = args[4];
|
||||
return true;
|
||||
} else {
|
||||
return true;
|
||||
return false;
|
||||
}
|
||||
} else {
|
||||
return false;
|
||||
}
|
||||
}
|
||||
|
||||
bool is_ite(expr const & e) {
|
||||
expr const & fn = get_app_fn(e);
|
||||
if (is_constant(fn) && const_name(fn) == get_ite_name()) {
|
||||
buffer<expr> args;
|
||||
get_app_args(e, args);
|
||||
if (args.size() == 5) return true;
|
||||
else return false;
|
||||
} else {
|
||||
return false;
|
||||
}
|
||||
}
|
||||
|
||||
bool is_iff(expr const & e) {
|
||||
expr const & fn = get_app_fn(e);
|
||||
return is_constant(fn) && const_name(fn) == get_iff_name();
|
||||
|
|
|
@ -111,7 +111,10 @@ expr instantiate_univ_param (expr const & e, name const & p, level const & l);
|
|||
|
||||
expr mk_true();
|
||||
expr mk_true_intro();
|
||||
|
||||
bool is_and(expr const & e, expr & arg1, expr & arg2);
|
||||
bool is_and(expr const & e);
|
||||
|
||||
expr mk_and(expr const & a, expr const & b);
|
||||
expr mk_and_intro(type_checker & tc, expr const & Ha, expr const & Hb);
|
||||
expr mk_and_elim_left(type_checker & tc, expr const & H);
|
||||
|
@ -149,6 +152,7 @@ bool is_or(expr const & e, expr & A, expr & B);
|
|||
/** \brief Return true if \c e is of the form <tt>(not arg)</tt>, and store \c arg in \c a.
|
||||
Return false otherwise */
|
||||
bool is_not(environment const & env, expr const & e, expr & a);
|
||||
bool is_not(environment const & env, expr const & e);
|
||||
expr mk_not(type_checker & tc, expr const & e);
|
||||
|
||||
/** \brief Create the term <tt>absurd e not_e : t</tt>. */
|
||||
|
@ -183,6 +187,7 @@ bool is_heq(expr const & e);
|
|||
bool is_heq(expr const & e, expr & A, expr & lhs, expr & B, expr & rhs);
|
||||
|
||||
bool is_ite(expr const & e, expr & c, expr & H, expr & A, expr & t, expr & f);
|
||||
bool is_ite(expr const & e);
|
||||
|
||||
bool is_iff(expr const & e);
|
||||
bool is_iff(expr const & e, expr & lhs, expr & rhs);
|
||||
|
|
|
@ -1,9 +1,11 @@
|
|||
set_option blast.simp false
|
||||
set_option blast.subst false
|
||||
|
||||
set_option blast.trace true
|
||||
example (a b c : Prop) : (a ↔ b) → ((a ∧ (c ∨ b)) ↔ (b ∧ (c ∨ a))) :=
|
||||
by blast
|
||||
|
||||
/-
|
||||
example (a b c : Prop) : (a ↔ b) → ((a ∧ (c ∨ (b ↔ a))) ↔ (b ∧ (c ∨ (a ↔ b)))) :=
|
||||
by blast
|
||||
|
||||
|
@ -14,3 +16,4 @@ definition lemma1 (a₁ a₂ b₁ b₂ : Prop) : (a₁ = b₁) → (a₂ ↔ b
|
|||
by blast
|
||||
|
||||
print lemma1
|
||||
-/
|
||||
|
|
13
tests/lean/run/blast_unit_preprocess.lean
Normal file
13
tests/lean/run/blast_unit_preprocess.lean
Normal file
|
@ -0,0 +1,13 @@
|
|||
-- Testing the unit pre-processor
|
||||
|
||||
open simplifier.unit_simp
|
||||
set_option simplify.max_steps 10000
|
||||
variables {A₁ A₂ A₃ A₄ B₁ B₂ B₃ B₄ : Prop}
|
||||
variables {A B C D E F G : Prop}
|
||||
variables {X : Type.{1}}
|
||||
|
||||
example (H : A ∨ B → E) (nE : ¬ E) : ¬ A ∧ ¬ B := by blast
|
||||
example (H : A ∨ B → E) (nE : ¬ E) : ¬ A ∧ ¬ B := by blast
|
||||
example (H : A ∨ B → E ∧ F) (nE : ¬ E) : ¬ A ∧ ¬ B := by blast
|
||||
example (H : A ∨ (B ∧ C) → E ∧ F) (c : C) (nE : ¬ E) : ¬ A ∧ ¬ B := by blast
|
||||
example (H : A ∨ (B ∧ C) → (G → E ∧ F)) (g : G) (c : C) (nE : ¬ E) : ¬ A ∧ ¬ B := by blast
|
|
@ -33,7 +33,7 @@ attribute if_ctx_simp_congr [congr]
|
|||
#simplify eq env 0 (ite b x y)
|
||||
end if_ctx_simp_congr
|
||||
|
||||
namespace if_congr_prop
|
||||
namespace if_ctx_congr_prop
|
||||
constants {b c x y u v : Prop} (dec_b : decidable b) (dec_c : decidable c)
|
||||
(h_c : b ↔ c) (h_t : c → (x ↔ u)) (h_e : ¬c → (y ↔ v))
|
||||
|
||||
|
@ -44,9 +44,8 @@ constants {b c x y u v : Prop} (dec_b : decidable b) (dec_c : decidable c)
|
|||
attribute h_e [simp]
|
||||
|
||||
attribute if_ctx_congr_prop [congr]
|
||||
|
||||
#simplify iff env 0 (ite b x y)
|
||||
end if_congr_prop
|
||||
end if_ctx_congr_prop
|
||||
|
||||
namespace if_ctx_simp_congr_prop
|
||||
constants (b c x y u v : Prop) (dec_b : decidable b)
|
||||
|
|
|
@ -11,6 +11,5 @@ attribute H3 [simp]
|
|||
attribute Hf [simp]
|
||||
attribute Hg [simp]
|
||||
|
||||
|
||||
#simplify iff env 2 (and P1 (and P2 P3))
|
||||
#simplify iff env 2 (and P1 (iff P2 P3))
|
||||
|
|
37
tests/lean/simplifier_unit_preprocess.lean
Normal file
37
tests/lean/simplifier_unit_preprocess.lean
Normal file
|
@ -0,0 +1,37 @@
|
|||
/-
|
||||
1. A ∧ B → C ==> { A → B → C }
|
||||
2. A ∨ B → C ==> { A → C, B → C }
|
||||
3. A → B ∧ C ==> { A → C, A → C }
|
||||
4. A ↔ B ==> { (A → B) ∧ (B → A) }
|
||||
5. Push negations inside ∧ and ∨ (when using classical)
|
||||
6. ite A B C ==> { A → B, ¬ A → C }
|
||||
-/
|
||||
|
||||
open simplifier.unit_simp
|
||||
set_option simplify.max_steps 100000
|
||||
|
||||
namespace vars
|
||||
variables {A B C D E F G : Prop}
|
||||
variables [A_dec : decidable A]
|
||||
|
||||
#simplify iff env 0 A → B ∧ C
|
||||
#simplify iff env 0 A → B ∧ C ∧ D
|
||||
#simplify iff env 0 A → B ∧ C → D
|
||||
#simplify iff env 0 A → B ∧ C → D ∧ E
|
||||
#simplify iff env 0 A → B ∧ C → D ∨ E → F ∧ G
|
||||
#simplify iff env 0 A → B ∧ C → D ∨ E → ((F → G) ∧ (G → F))
|
||||
|
||||
end vars
|
||||
|
||||
namespace cs
|
||||
|
||||
constants {A B C D E F G : Prop}
|
||||
constants [A_dec : decidable A]
|
||||
|
||||
attribute A_dec [instance]
|
||||
|
||||
#simplify iff env 0 ite A B C
|
||||
#simplify iff env 0 ite A B C → B ∧ C
|
||||
#simplify iff env 0 A → (ite A B C)
|
||||
#simplify iff env 0 A → B ∧ (ite A B C) → D ∨ E → ((F → G) ∧ (G → F))
|
||||
end cs
|
13
tests/lean/simplifier_unit_preprocess.lean.expected.out
Normal file
13
tests/lean/simplifier_unit_preprocess.lean.expected.out
Normal file
|
@ -0,0 +1,13 @@
|
|||
(A → B) ∧ (A → C)
|
||||
(A → B) ∧ (A → C) ∧ (A → D)
|
||||
A → B → C → D
|
||||
(A → B → C → D) ∧ (A → B → C → E)
|
||||
(A → B → C → D → F) ∧ (A → B → C → D → G) ∧ (A → B → C → E → F) ∧ (A →
|
||||
B → C → E → G)
|
||||
(A → B → C → D → F → G) ∧ (A → B → C → D → G → F) ∧ (A → B → C → E → F → G) ∧ (A →
|
||||
B → C → E → G → F)
|
||||
(A → B) ∧ (¬A → C)
|
||||
((A → B) → (¬A → C) → B) ∧ ((A → B) → (¬A → C) → C)
|
||||
A → B
|
||||
(A → B → D → F → G) ∧ (A → B → D → G → F) ∧ (A → B → E → F → G) ∧ (A →
|
||||
B → E → G → F)
|
Loading…
Reference in a new issue