feat(library/simplifier): improve simplification by evaluation
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
08053c1172
commit
69d7ee316f
7 changed files with 198 additions and 6 deletions
|
@ -298,6 +298,12 @@ theorem imp_falser (a : Bool) : (a → false) ↔ ¬ a
|
|||
theorem imp_falsel (a : Bool) : (false → a) ↔ true
|
||||
:= case (λ x, (false → x) ↔ true) trivial trivial a
|
||||
|
||||
theorem not_true : ¬ true ↔ false
|
||||
:= trivial
|
||||
|
||||
theorem not_false : ¬ false ↔ true
|
||||
:= trivial
|
||||
|
||||
theorem not_and (a b : Bool) : ¬ (a ∧ b) ↔ ¬ a ∨ ¬ b
|
||||
:= case (λ x, ¬ (x ∧ b) ↔ ¬ x ∨ ¬ b)
|
||||
(case (λ y, ¬ (true ∧ y) ↔ ¬ true ∨ ¬ y) trivial trivial b)
|
||||
|
|
Binary file not shown.
|
@ -86,6 +86,8 @@ MK_CONSTANT(imp_truer_fn, name("imp_truer"));
|
|||
MK_CONSTANT(imp_truel_fn, name("imp_truel"));
|
||||
MK_CONSTANT(imp_falser_fn, name("imp_falser"));
|
||||
MK_CONSTANT(imp_falsel_fn, name("imp_falsel"));
|
||||
MK_CONSTANT(not_true, name("not_true"));
|
||||
MK_CONSTANT(not_false, name("not_false"));
|
||||
MK_CONSTANT(not_and_fn, name("not_and"));
|
||||
MK_CONSTANT(not_and_elim_fn, name("not_and_elim"));
|
||||
MK_CONSTANT(not_or_fn, name("not_or"));
|
||||
|
|
|
@ -249,6 +249,10 @@ inline expr mk_imp_falser_th(expr const & e1) { return mk_app({mk_imp_falser_fn(
|
|||
expr mk_imp_falsel_fn();
|
||||
bool is_imp_falsel_fn(expr const & e);
|
||||
inline expr mk_imp_falsel_th(expr const & e1) { return mk_app({mk_imp_falsel_fn(), e1}); }
|
||||
expr mk_not_true();
|
||||
bool is_not_true(expr const & e);
|
||||
expr mk_not_false();
|
||||
bool is_not_false(expr const & e);
|
||||
expr mk_not_and_fn();
|
||||
bool is_not_and_fn(expr const & e);
|
||||
inline expr mk_not_and_th(expr const & e1, expr const & e2) { return mk_app({mk_not_and_fn(), e1, e2}); }
|
||||
|
|
|
@ -11,6 +11,7 @@ Author: Leonardo de Moura
|
|||
#include "kernel/type_checker.h"
|
||||
#include "kernel/free_vars.h"
|
||||
#include "kernel/instantiate.h"
|
||||
#include "kernel/normalizer.h"
|
||||
#include "kernel/kernel.h"
|
||||
#include "library/heq_decls.h"
|
||||
#include "library/kernel_bindings.h"
|
||||
|
@ -168,6 +169,11 @@ class simplifier_fn {
|
|||
return m_tc.ensure_pi(e, m_ctx);
|
||||
}
|
||||
|
||||
expr normalize(expr const & e) {
|
||||
normalizer & proc = m_tc.get_normalizer();
|
||||
return proc(e, m_ctx, true);
|
||||
}
|
||||
|
||||
expr mk_congr1_th(expr const & f_type, expr const & f, expr const & new_f, expr const & a, expr const & Heq_f) {
|
||||
expr const & A = abst_domain(f_type);
|
||||
expr B = lower_free_vars(abst_body(f_type), 1, 1);
|
||||
|
@ -392,6 +398,38 @@ class simplifier_fn {
|
|||
}
|
||||
}
|
||||
|
||||
/** \brief Return true when \c e is a value from the point of view of the simplifier */
|
||||
static bool is_value(expr const & e) {
|
||||
// Currently only semantic attachments are treated as value.
|
||||
// We may relax that in the future.
|
||||
return ::lean::is_value(e);
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Return true iff the simplifier should use the evaluator/normalizer to reduce application
|
||||
*/
|
||||
bool evaluate_app(expr const & e) const {
|
||||
lean_assert(is_app(e));
|
||||
// only evaluate if it is enabled
|
||||
if (!m_eval)
|
||||
return false;
|
||||
// if all arguments are values, we should evaluate
|
||||
if (std::all_of(args(e).begin()+1, args(e).end(), [](expr const & a) { return is_value(a); }))
|
||||
return true;
|
||||
// The previous test fails for equality/disequality because the first arguments are types.
|
||||
// Should we have something more general for cases like that?
|
||||
// Some possibilities:
|
||||
// - We have a table mapping constants to argument positions. The positions tell the simplifier
|
||||
// which arguments must be value to trigger evaluation.
|
||||
// - We have an external predicate that is invoked by the simplifier to decide whether to normalize/evaluate an
|
||||
// expression.
|
||||
unsigned num = num_args(e);
|
||||
return
|
||||
(is_eq(e) || is_neq(e) || is_heq(e)) &&
|
||||
is_value(arg(e, num-2)) &&
|
||||
is_value(arg(e, num-1));
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Given (applications) lhs and rhs s.t. lhs = rhs.m_out
|
||||
with proof rhs.m_proof, this method applies rewrite rules, beta
|
||||
|
@ -404,16 +442,17 @@ class simplifier_fn {
|
|||
result rewrite_app(expr const & lhs, result const & rhs) {
|
||||
lean_assert(is_app(rhs.m_out));
|
||||
lean_assert(is_app(lhs));
|
||||
expr f = arg(rhs.m_out, 0);
|
||||
if (m_eval && is_value(f)) {
|
||||
optional<expr> new_rhs = to_value(f).normalize(num_args(rhs.m_out), &arg(rhs.m_out, 0));
|
||||
if (new_rhs) {
|
||||
if (evaluate_app(rhs.m_out)) {
|
||||
// try to evaluate if all arguments are values.
|
||||
expr new_rhs = normalize(rhs.m_out);
|
||||
if (is_value(new_rhs)) {
|
||||
// We don't need to create a new proof term since rhs.m_out and new_rhs are
|
||||
// definitionally equal.
|
||||
return rewrite(lhs, result(*new_rhs, rhs.m_proof, rhs.m_heq_proof));
|
||||
return rewrite(lhs, result(new_rhs, rhs.m_proof, rhs.m_heq_proof));
|
||||
}
|
||||
}
|
||||
|
||||
expr f = arg(rhs.m_out, 0);
|
||||
if (m_beta && is_lambda(f)) {
|
||||
expr new_rhs = head_beta_reduce(rhs.m_out);
|
||||
// rhs.m_out and new_rhs are also definitionally equal
|
||||
|
@ -492,7 +531,12 @@ class simplifier_fn {
|
|||
}
|
||||
}
|
||||
}
|
||||
return rhs;
|
||||
if (!m_single_pass && lhs != rhs.m_out) {
|
||||
result new_rhs = simplify(rhs.m_out);
|
||||
return mk_trans_result(lhs, rhs, new_rhs);
|
||||
} else {
|
||||
return rhs;
|
||||
}
|
||||
}
|
||||
|
||||
result simplify_var(expr const & e) {
|
||||
|
@ -604,6 +648,7 @@ class simplifier_fn {
|
|||
|
||||
result simplify_pi(expr const & e) {
|
||||
lean_assert(is_pi(e));
|
||||
// TODO(Leo): handle implication, i.e., e is_proposition and is_arrow
|
||||
if (m_has_heq) {
|
||||
// TODO(Leo)
|
||||
return result(e);
|
||||
|
|
100
tests/lean/simp3.lean
Normal file
100
tests/lean/simp3.lean
Normal file
|
@ -0,0 +1,100 @@
|
|||
definition double x := x + x
|
||||
|
||||
(*
|
||||
function show(x) print(x) end
|
||||
local t1 = parse_lean("0 ≠ 1")
|
||||
show(simplify(t1))
|
||||
local t2 = parse_lean("3 ≥ 2")
|
||||
show(simplify(t2))
|
||||
local t3 = parse_lean("double (double 2) + 1")
|
||||
show(simplify(t3))
|
||||
local t4 = parse_lean("0 = 1")
|
||||
show(simplify(t4))
|
||||
*)
|
||||
|
||||
(*
|
||||
local opt = options({"simplifier", "unfold"}, true, {"simplifier", "eval"}, false)
|
||||
local t1 = parse_lean("double (double 2) + 1 ≥ 3")
|
||||
show(simplify(t1, 'default', get_environment(), context(), opt))
|
||||
*)
|
||||
|
||||
set_opaque Nat::ge false
|
||||
add_rewrite Nat::add_assoc
|
||||
add_rewrite Nat::distributer
|
||||
add_rewrite Nat::distributel
|
||||
|
||||
(*
|
||||
local opt = options({"simplifier", "unfold"}, true, {"simplifier", "eval"}, false)
|
||||
local t1 = parse_lean("2 * double (double 2) + 1 ≥ 3")
|
||||
show(simplify(t1, 'default', get_environment(), context(), opt))
|
||||
*)
|
||||
|
||||
variables a b c d : Nat
|
||||
|
||||
import if_then_else
|
||||
add_rewrite if_true
|
||||
add_rewrite if_false
|
||||
add_rewrite if_a_a
|
||||
add_rewrite Nat::add_zeror
|
||||
add_rewrite Nat::add_zerol
|
||||
add_rewrite eq_id
|
||||
|
||||
(*
|
||||
local t1 = parse_lean("(a + b) * (c + d)")
|
||||
local r, pr = simplify(t1)
|
||||
print(r)
|
||||
print(pr)
|
||||
*)
|
||||
|
||||
theorem congr2_congr1 {A B C : TypeU} {f g : A → B} (h : B → C) (Hfg : f = g) (a : A) :
|
||||
congr2 h (congr1 a Hfg) = congr2 (λ x, h (x a)) Hfg
|
||||
:= proof_irrel (congr2 h (congr1 a Hfg)) (congr2 (λ x, h (x a)) Hfg)
|
||||
|
||||
theorem congr2_congr2 {A B C : TypeU} {a b : A} (f : A → B) (h : B → C) (Hab : a = b) :
|
||||
congr2 h (congr2 f Hab) = congr2 (λ x, h (f x)) Hab
|
||||
:= proof_irrel (congr2 h (congr2 f Hab)) (congr2 (λ x, h (f x)) Hab)
|
||||
|
||||
theorem congr1_congr2 {A B C : TypeU} {a b : A} (f : A → B → C) (Hab : a = b) (c : B):
|
||||
congr1 c (congr2 f Hab) = congr2 (λ x, f x c) Hab
|
||||
:= proof_irrel (congr1 c (congr2 f Hab)) (congr2 (λ x, f x c) Hab)
|
||||
|
||||
rewrite_set proofsimp
|
||||
add_rewrite congr2_congr1 congr2_congr2 congr1_congr2 : proofsimp
|
||||
|
||||
(*
|
||||
local t2 = parse_lean("(if a > 0 then b else b + 0) + 10 = (if a > 0 then b else b) + 10")
|
||||
local r, pr = simplify(t2)
|
||||
print(r)
|
||||
print(pr)
|
||||
show(simplify(pr, 'proofsimp'))
|
||||
*)
|
||||
|
||||
|
||||
(*
|
||||
local t1 = parse_lean("(a + b) * (a + b)")
|
||||
local t2 = simplify(t1)
|
||||
print(t2)
|
||||
*)
|
||||
|
||||
-- add_rewrite imp_truer imp_truel imp_falsel imp_falser not_true not_false
|
||||
-- print rewrite_set
|
||||
|
||||
(*
|
||||
local t1 = parse_lean("true → false")
|
||||
print(simplify(t1))
|
||||
*)
|
||||
|
||||
(*
|
||||
local t1 = parse_lean("true → true")
|
||||
print(simplify(t1))
|
||||
*)
|
||||
|
||||
(*
|
||||
local t1 = parse_lean("false → false")
|
||||
print(simplify(t1))
|
||||
*)
|
||||
|
||||
(*
|
||||
local t1 = parse_lean("true ↔ false")
|
||||
print(simplify(t1))
|
||||
*)
|
35
tests/lean/simp3.lean.expected.out
Normal file
35
tests/lean/simp3.lean.expected.out
Normal file
|
@ -0,0 +1,35 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Defined: double
|
||||
⊤
|
||||
⊤
|
||||
9
|
||||
⊥
|
||||
2 + 2 + (2 + 2) + 1 ≥ 3
|
||||
3 ≤ 2 * 2 + 2 * 2 + 2 * 2 + 2 * 2 + 1
|
||||
Assumed: a
|
||||
Assumed: b
|
||||
Assumed: c
|
||||
Assumed: d
|
||||
Imported 'if_then_else'
|
||||
a * c + a * d + b * c + b * d
|
||||
trans (Nat::distributel a b (c + d))
|
||||
(trans (congr (congr2 Nat::add (Nat::distributer a c d)) (Nat::distributer b c d))
|
||||
(Nat::add_assoc (a * c + a * d) (b * c) (b * d)))
|
||||
Proved: congr2_congr1
|
||||
Proved: congr2_congr2
|
||||
Proved: congr1_congr2
|
||||
⊤
|
||||
trans (congr (congr2 eq
|
||||
(congr1 10
|
||||
(congr2 Nat::add (trans (congr2 (ite (a > 0) b) (Nat::add_zeror b)) (if_a_a (a > 0) b)))))
|
||||
(congr1 10 (congr2 Nat::add (if_a_a (a > 0) b))))
|
||||
(eq_id (b + 10))
|
||||
let κ::1 := congr2 (λ x : ℕ → ℕ, eq (x 10))
|
||||
(congr2 Nat::add (trans (congr2 (ite (a > 0) b) (Nat::add_zeror b)) (if_a_a (a > 0) b)))
|
||||
in trans (congr κ::1 (congr1 10 (congr2 Nat::add (if_a_a (a > 0) b)))) (eq_id (b + 10))
|
||||
a * a + a * b + b * a + b * b
|
||||
⊤ → ⊥ refl (⊤ → ⊥)
|
||||
⊤ → ⊤ refl (⊤ → ⊤)
|
||||
⊥ → ⊥ refl (⊥ → ⊥)
|
||||
⊥ refl ⊥
|
Loading…
Reference in a new issue