feat(builtin/kernel): create default rule set in the kernel, and adjust unit tests
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
3bbadddc94
commit
bbf6e6a256
15 changed files with 28 additions and 4 deletions
|
@ -69,6 +69,9 @@ theorem not_intro {a : Bool} (H : a → false) : ¬ a
|
|||
theorem eta {A : TypeU} {B : A → TypeU} (f : ∀ x : A, B x) : (λ x : A, f x) = f
|
||||
:= funext (λ x : A, refl (f x))
|
||||
|
||||
-- create default rewrite rule set
|
||||
(* mk_rewrite_rule_set() *)
|
||||
|
||||
theorem trivial : true
|
||||
:= refl true
|
||||
|
||||
|
|
Binary file not shown.
|
@ -133,9 +133,9 @@ class simplifier_fn {
|
|||
match_fn m_match_fn;
|
||||
|
||||
struct result {
|
||||
expr m_out;
|
||||
optional<expr> m_proof;
|
||||
bool m_heq_proof;
|
||||
expr m_out; // the result of a simplification step
|
||||
optional<expr> m_proof; // a proof that the result is equal to the input (when m_proofs_enabled)
|
||||
bool m_heq_proof; // true if the proof is for heterogeneous equality
|
||||
explicit result(expr const & out, bool heq_proof = false):
|
||||
m_out(out), m_heq_proof(heq_proof) {}
|
||||
result(expr const & out, expr const & pr, bool heq_proof = false):
|
||||
|
|
|
@ -16,6 +16,7 @@ Author: Leonardo de Moura
|
|||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lean/operator_info.h"
|
||||
#include "frontends/lean/pp.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
static void tst1() {
|
||||
|
@ -208,6 +209,7 @@ static void tst11() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
tst2();
|
||||
tst3();
|
||||
|
|
|
@ -16,6 +16,7 @@ Author: Leonardo de Moura
|
|||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lean/pp.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
static void parse(environment const & env, io_state const & ios, char const * str) {
|
||||
|
@ -109,6 +110,7 @@ static void tst3() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
tst2();
|
||||
tst3();
|
||||
|
|
|
@ -11,6 +11,7 @@ Author: Leonardo de Moura
|
|||
#include "library/io_state_stream.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lean/pp.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
static expr mk_shared_expr(unsigned depth) {
|
||||
|
@ -105,6 +106,7 @@ static void tst6() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
tst2();
|
||||
tst3();
|
||||
|
|
|
@ -17,6 +17,7 @@ Author: Leonardo de Moura
|
|||
#include "library/io_state_stream.h"
|
||||
#include "library/arith/arith.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
static void tst1() {
|
||||
|
@ -272,6 +273,7 @@ static void tst13() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
enable_trace("is_convertible");
|
||||
tst1();
|
||||
tst2();
|
||||
|
|
|
@ -24,6 +24,7 @@ Author: Leonardo de Moura
|
|||
#include "library/placeholder.h"
|
||||
#include "library/arith/arith.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
static std::ostream & operator<<(std::ostream & out, metavar_env const & menv) {
|
||||
|
@ -631,6 +632,7 @@ static void tst28() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
tst2();
|
||||
tst3();
|
||||
|
|
|
@ -22,6 +22,7 @@ Author: Leonardo de Moura
|
|||
#include "library/deep_copy.h"
|
||||
#include "library/arith/int.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
expr normalize(expr const & e) {
|
||||
|
@ -344,6 +345,7 @@ static void tst11() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst_church_numbers();
|
||||
tst1();
|
||||
tst2();
|
||||
|
|
|
@ -16,6 +16,7 @@ Author: Leonardo de Moura
|
|||
#include "library/deep_copy.h"
|
||||
#include "library/arith/arith.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
expr norm(expr const & e, environment & env) {
|
||||
|
@ -65,6 +66,7 @@ static void tst1() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
return 0;
|
||||
}
|
||||
|
|
|
@ -24,6 +24,7 @@ Author: Leonardo de Moura
|
|||
#include "library/io_state_stream.h"
|
||||
#include "library/arith/arith.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
expr c(char const * n) { return mk_constant(n); }
|
||||
|
@ -488,6 +489,7 @@ static void tst21() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
tst2();
|
||||
tst3();
|
||||
|
|
|
@ -13,6 +13,7 @@ Author: Leonardo de Moura
|
|||
#include "library/io_state_stream.h"
|
||||
#include "library/arith/arith.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
static void tst0() {
|
||||
|
@ -132,6 +133,7 @@ static void tst6() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst0();
|
||||
tst1();
|
||||
tst2();
|
||||
|
|
|
@ -18,6 +18,7 @@ Author: Leonardo de Moura
|
|||
#include "library/arith/arith.h"
|
||||
#include "library/elaborator/elaborator.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
static void tst1() {
|
||||
|
@ -867,6 +868,7 @@ void tst27() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
tst2();
|
||||
tst3();
|
||||
|
|
|
@ -17,6 +17,7 @@ Author: Leonardo de Moura
|
|||
#include "library/tactic/tactic.h"
|
||||
#include "library/tactic/boolean_tactics.h"
|
||||
#include "frontends/lean/frontend.h"
|
||||
#include "frontends/lua/register_modules.h"
|
||||
using namespace lean;
|
||||
|
||||
tactic loop_tactic() {
|
||||
|
@ -116,6 +117,7 @@ static void tst1() {
|
|||
|
||||
int main() {
|
||||
save_stack_info();
|
||||
register_modules();
|
||||
tst1();
|
||||
return has_violations() ? 1 : 0;
|
||||
}
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
mk_rewrite_rule_set()
|
||||
add_rewrite_rules({"Nat", "add_zerol"})
|
||||
add_rewrite_rules({"Nat", "add_zeror"})
|
||||
parse_lean_cmds([[
|
||||
|
|
Loading…
Reference in a new issue