feat(library/tactic): add simplify_tactic based on the simplifier
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
5e6c1d4904
commit
4d25cb7f47
21 changed files with 236 additions and 4 deletions
|
@ -173,6 +173,9 @@ theorem refute {a : Bool} (H : ¬ a → false) : a
|
|||
theorem symm {A : TypeU} {a b : A} (H : a = b) : b = a
|
||||
:= subst (refl a) H
|
||||
|
||||
theorem eqmpr {a b : Bool} (H1 : a = b) (H2 : b) : a
|
||||
:= (symm H1) ◂ H2
|
||||
|
||||
theorem trans {A : TypeU} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c
|
||||
:= subst H1 H2
|
||||
|
||||
|
|
Binary file not shown.
|
@ -16,3 +16,9 @@ const_tactic("unfold_all", unfold_tac)
|
|||
const_tactic("beta", beta_tac)
|
||||
tactic_macro("apply", { macro_arg.Expr }, function (env, e) return apply_tac(e) end)
|
||||
tactic_macro("unfold", { macro_arg.Id }, function (env, id) return unfold_tac(id) end)
|
||||
|
||||
tactic_macro("simp", { macro_arg.Ids },
|
||||
function (env, ids)
|
||||
return simp_tac(ids)
|
||||
end
|
||||
)
|
||||
|
|
|
@ -372,6 +372,15 @@ parser_imp::macro_result parser_imp::parse_macro(list<macro_arg_kind> const & ar
|
|||
args.emplace_back(k, &n);
|
||||
return parse_macro(tail(arg_kinds), fn, prec, args, p);
|
||||
}
|
||||
case macro_arg_kind::Ids: {
|
||||
buffer<name> ns;
|
||||
args.emplace_back(k, &ns);
|
||||
while (curr_is_identifier()) {
|
||||
ns.push_back(curr_name());
|
||||
next();
|
||||
}
|
||||
return parse_macro(tail(arg_kinds), fn, prec, args, p);
|
||||
}
|
||||
case macro_arg_kind::Tactic: {
|
||||
tactic t = parse_tactic_expr();
|
||||
args.emplace_back(k, &t);
|
||||
|
@ -428,6 +437,17 @@ parser_imp::macro_result parser_imp::parse_macro(list<macro_arg_kind> const & ar
|
|||
case macro_arg_kind::Id:
|
||||
push_name(L, *static_cast<name*>(arg));
|
||||
break;
|
||||
case macro_arg_kind::Ids: {
|
||||
buffer<name> const & ids = *static_cast<buffer<name>*>(arg);
|
||||
lua_newtable(L);
|
||||
int i = 1;
|
||||
for (auto const & id : ids) {
|
||||
push_name(L, id);
|
||||
lua_rawseti(L, -2, i);
|
||||
i = i + 1;
|
||||
}
|
||||
break;
|
||||
}
|
||||
case macro_arg_kind::String:
|
||||
lua_pushstring(L, static_cast<std::string*>(arg)->c_str());
|
||||
break;
|
||||
|
|
|
@ -131,6 +131,7 @@ void open_macros(lua_State * L) {
|
|||
SET_ENUM("Exprs", macro_arg_kind::Exprs);
|
||||
SET_ENUM("Parameters", macro_arg_kind::Parameters);
|
||||
SET_ENUM("Id", macro_arg_kind::Id);
|
||||
SET_ENUM("Ids", macro_arg_kind::Ids);
|
||||
SET_ENUM("String", macro_arg_kind::String);
|
||||
SET_ENUM("Int", macro_arg_kind::Int);
|
||||
SET_ENUM("Comma", macro_arg_kind::Comma);
|
||||
|
|
|
@ -26,7 +26,7 @@ struct parameter {
|
|||
};
|
||||
typedef buffer<parameter> parameter_buffer;
|
||||
|
||||
enum class macro_arg_kind { Expr, Exprs, Parameters, Id, Int, String, Comma, Assign, Tactic, Tactics };
|
||||
enum class macro_arg_kind { Expr, Exprs, Parameters, Id, Ids, Int, String, Comma, Assign, Tactic, Tactics };
|
||||
struct macro {
|
||||
list<macro_arg_kind> m_arg_kinds;
|
||||
luaref m_fn;
|
||||
|
|
|
@ -54,6 +54,7 @@ MK_CONSTANT(or_intror_fn, name("or_intror"));
|
|||
MK_CONSTANT(or_elim_fn, name("or_elim"));
|
||||
MK_CONSTANT(refute_fn, name("refute"));
|
||||
MK_CONSTANT(symm_fn, name("symm"));
|
||||
MK_CONSTANT(eqmpr_fn, name("eqmpr"));
|
||||
MK_CONSTANT(trans_fn, name("trans"));
|
||||
MK_CONSTANT(ne_symm_fn, name("ne_symm"));
|
||||
MK_CONSTANT(eq_ne_trans_fn, name("eq_ne_trans"));
|
||||
|
|
|
@ -155,6 +155,9 @@ inline expr mk_refute_th(expr const & e1, expr const & e2) { return mk_app({mk_r
|
|||
expr mk_symm_fn();
|
||||
bool is_symm_fn(expr const & e);
|
||||
inline expr mk_symm_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4) { return mk_app({mk_symm_fn(), e1, e2, e3, e4}); }
|
||||
expr mk_eqmpr_fn();
|
||||
bool is_eqmpr_fn(expr const & e);
|
||||
inline expr mk_eqmpr_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4) { return mk_app({mk_eqmpr_fn(), e1, e2, e3, e4}); }
|
||||
expr mk_trans_fn();
|
||||
bool is_trans_fn(expr const & e);
|
||||
inline expr mk_trans_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4, expr const & e5, expr const & e6) { return mk_app({mk_trans_fn(), e1, e2, e3, e4, e5, e6}); }
|
||||
|
|
|
@ -424,6 +424,10 @@ context metavar_env_cell::instantiate_metavars(context const & ctx) const {
|
|||
return context(new_entries.size(), new_entries.data());
|
||||
}
|
||||
|
||||
metavar_env::metavar_env(ro_metavar_env const & s):m_ptr(s.m_ptr) {
|
||||
if (m_ptr) m_ptr->inc_ref();
|
||||
}
|
||||
|
||||
bool cached_metavar_env::update(optional<metavar_env> const & menv) {
|
||||
if (!menv) {
|
||||
m_menv = none_menv();
|
||||
|
|
|
@ -196,6 +196,7 @@ public:
|
|||
context instantiate_metavars(context const & ctx) const;
|
||||
};
|
||||
|
||||
class ro_metavar_env;
|
||||
/**
|
||||
\brief Reference to metavariable environment (cell).
|
||||
*/
|
||||
|
@ -205,6 +206,8 @@ class metavar_env {
|
|||
friend class metavar_env_cell;
|
||||
metavar_env_cell * m_ptr;
|
||||
explicit metavar_env(metavar_env_cell * ptr):m_ptr(ptr) { if (m_ptr) m_ptr->inc_ref(); }
|
||||
friend class type_checker;
|
||||
explicit metavar_env(ro_metavar_env const & s);
|
||||
public:
|
||||
metavar_env():m_ptr(new metavar_env_cell()) { m_ptr->inc_ref(); }
|
||||
metavar_env(name const & prefix):m_ptr(new metavar_env_cell(prefix)) { m_ptr->inc_ref(); }
|
||||
|
@ -243,6 +246,7 @@ inline context instantiate_metavars(optional<metavar_env> const & menv, context
|
|||
\brief Read-only reference to metavariable environment (cell).
|
||||
*/
|
||||
class ro_metavar_env {
|
||||
friend class metavar_env;
|
||||
metavar_env_cell * m_ptr;
|
||||
public:
|
||||
ro_metavar_env():m_ptr(new metavar_env_cell()) { m_ptr->inc_ref(); }
|
||||
|
|
|
@ -482,6 +482,10 @@ expr type_checker::infer_type(expr const & e, context const & ctx, metavar_env c
|
|||
expr type_checker::infer_type(expr const & e, context const & ctx, metavar_env const & menv) {
|
||||
return m_ptr->infer_type(e, ctx, some_menv(menv), nullptr);
|
||||
}
|
||||
expr type_checker::infer_type(expr const & e, context const & ctx, ro_metavar_env const & menv) {
|
||||
// metavariable environment is not updated when unification constraints are not provided
|
||||
return infer_type(e, ctx, metavar_env(menv));
|
||||
}
|
||||
expr type_checker::infer_type(expr const & e, context const & ctx) {
|
||||
return infer_type(e, ctx, none_menv(), nullptr);
|
||||
}
|
||||
|
@ -518,6 +522,10 @@ bool type_checker::is_proposition(expr const & e, context const & ctx) {
|
|||
bool type_checker::is_proposition(expr const & e, context const & ctx, metavar_env const & menv) {
|
||||
return is_proposition(e, ctx, some_menv(menv));
|
||||
}
|
||||
bool type_checker::is_proposition(expr const & e, context const & ctx, ro_metavar_env const & menv) {
|
||||
// metavariable environment is not updated when unification constraints are not provided
|
||||
return is_proposition(e, ctx, metavar_env(menv));
|
||||
}
|
||||
bool type_checker::is_flex_proposition(expr e, context ctx, optional<metavar_env> const & menv) {
|
||||
while (is_pi(e)) {
|
||||
ctx = extend(ctx, abst_name(e), abst_domain(e));
|
||||
|
|
|
@ -50,6 +50,7 @@ public:
|
|||
expr infer_type(expr const & e, context const & ctx, optional<metavar_env> const & menv, buffer<unification_constraint> * new_constraints);
|
||||
expr infer_type(expr const & e, context const & ctx, metavar_env const & menv, buffer<unification_constraint> & new_constraints);
|
||||
expr infer_type(expr const & e, context const & ctx, metavar_env const & menv);
|
||||
expr infer_type(expr const & e, context const & ctx, ro_metavar_env const & menv);
|
||||
|
||||
/**
|
||||
\brief Return the type of \c e in the context \c ctx.
|
||||
|
@ -95,6 +96,7 @@ public:
|
|||
/** \brief Return true iff \c e is a proposition (i.e., it has type Bool) */
|
||||
bool is_proposition(expr const & e, context const & ctx, optional<metavar_env> const & menv);
|
||||
bool is_proposition(expr const & e, context const & ctx, metavar_env const & menv);
|
||||
bool is_proposition(expr const & e, context const & ctx, ro_metavar_env const & menv);
|
||||
bool is_proposition(expr const & e, context const & ctx = context());
|
||||
|
||||
/** \brief Return true iff \c e is a proposition or is a Pi s.t. the range is a flex_proposition */
|
||||
|
|
|
@ -8,6 +8,7 @@ Author: Leonardo de Moura
|
|||
#include "util/lua.h"
|
||||
#include "kernel/environment.h"
|
||||
#include "library/expr_pair.h"
|
||||
#include "library/simplifier/rewrite_rule_set.h"
|
||||
|
||||
namespace lean {
|
||||
expr_pair simplify(expr const & e, ro_environment const & env, context const & ctx, options const & pts,
|
||||
|
|
|
@ -1,4 +1,5 @@
|
|||
add_library(tactic goal.cpp proof_builder.cpp cex_builder.cpp
|
||||
proof_state.cpp tactic.cpp boolean_tactics.cpp apply_tactic.cpp)
|
||||
proof_state.cpp tactic.cpp boolean_tactics.cpp apply_tactic.cpp
|
||||
simplify_tactic.cpp)
|
||||
|
||||
target_link_libraries(tactic ${LEAN_LIBS})
|
||||
|
|
|
@ -13,6 +13,7 @@ Author: Leonardo de Moura
|
|||
#include "library/tactic/tactic.h"
|
||||
#include "library/tactic/boolean_tactics.h"
|
||||
#include "library/tactic/apply_tactic.h"
|
||||
#include "library/tactic/simplify_tactic.h"
|
||||
|
||||
namespace lean {
|
||||
inline void open_tactic_module(lua_State * L) {
|
||||
|
@ -23,6 +24,7 @@ inline void open_tactic_module(lua_State * L) {
|
|||
open_tactic(L);
|
||||
open_boolean_tactics(L);
|
||||
open_apply_tactic(L);
|
||||
open_simplify_tactic(L);
|
||||
}
|
||||
inline void register_tactic_module() {
|
||||
script_state::register_module(open_tactic_module);
|
||||
|
|
116
src/library/tactic/simplify_tactic.cpp
Normal file
116
src/library/tactic/simplify_tactic.cpp
Normal file
|
@ -0,0 +1,116 @@
|
|||
/*
|
||||
Copyright (c) 2013 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#include <vector>
|
||||
#include "util/buffer.h"
|
||||
#include "util/sexpr/option_declarations.h"
|
||||
#include "kernel/type_checker.h"
|
||||
#include "kernel/kernel.h"
|
||||
#include "library/simplifier/simplifier.h"
|
||||
#include "library/tactic/tactic.h"
|
||||
|
||||
#ifndef LEAN_SIMP_TAC_ASSUMPTIONS
|
||||
#define LEAN_SIMP_TAC_ASSUMPTIONS true
|
||||
#endif
|
||||
|
||||
namespace lean {
|
||||
static name g_simp_tac_assumptions {"simp_tac", "assumptions"};
|
||||
RegisterBoolOption(g_simp_tac_assumptions, LEAN_SIMP_TAC_ASSUMPTIONS, "(simplifier tactic) use goal assumptions as rewrite rules");
|
||||
bool get_simp_tac_assumptions(options const & opts) { return opts.get_bool(g_simp_tac_assumptions, LEAN_SIMP_TAC_ASSUMPTIONS); }
|
||||
|
||||
static name g_assumption("assump");
|
||||
|
||||
static optional<proof_state> simplify_tactic(ro_environment const & env, io_state const & ios, proof_state const & s,
|
||||
unsigned num_ns, name const * ns, options const & extra_opts) {
|
||||
if (empty(s.get_goals()))
|
||||
return none_proof_state();
|
||||
options opts = join(extra_opts, ios.get_options());
|
||||
|
||||
auto const & p = head(s.get_goals());
|
||||
name const & gname = p.first;
|
||||
goal const & g = p.second;
|
||||
ro_metavar_env const & menv = s.get_menv();
|
||||
type_checker tc(env);
|
||||
|
||||
bool use_assumptions = get_simp_tac_assumptions(opts);
|
||||
buffer<rewrite_rule_set> rule_sets;
|
||||
if (use_assumptions) {
|
||||
rule_sets.push_back(rewrite_rule_set(env));
|
||||
rewrite_rule_set & rs = rule_sets[0];
|
||||
for (auto const & p : g.get_hypotheses()) {
|
||||
if (tc.is_proposition(p.second, context(), menv)) {
|
||||
expr H = mk_constant(p.first, p.second);
|
||||
rs.insert(g_assumption, p.second, H);
|
||||
}
|
||||
}
|
||||
}
|
||||
for (unsigned i = 0; i < num_ns; i++) {
|
||||
rule_sets.push_back(get_rewrite_rule_set(env, ns[i]));
|
||||
}
|
||||
|
||||
expr conclusion = g.get_conclusion();
|
||||
auto r = simplify(conclusion, env, context(), opts, rule_sets.size(), rule_sets.data());
|
||||
expr new_conclusion = r.first;
|
||||
expr eq_proof = r.second;
|
||||
if (new_conclusion == g.get_conclusion())
|
||||
return optional<proof_state>(s);
|
||||
bool solved = is_true(new_conclusion);
|
||||
goals rest_gs = tail(s.get_goals());
|
||||
goals new_gs;
|
||||
if (solved)
|
||||
new_gs = rest_gs;
|
||||
else
|
||||
new_gs = goals(mk_pair(gname, update(g, new_conclusion)), rest_gs);
|
||||
proof_builder pb = s.get_proof_builder();
|
||||
proof_builder new_pb = mk_proof_builder([=](proof_map const & m, assignment const & a) -> expr {
|
||||
proof_map new_m(m);
|
||||
if (solved)
|
||||
new_m.insert(gname, mk_eqt_elim_th(conclusion, eq_proof));
|
||||
else
|
||||
new_m.insert(gname, mk_eqmpr_th(conclusion, new_conclusion, eq_proof, find(m, gname)));
|
||||
return pb(new_m, a);
|
||||
});
|
||||
return some(proof_state(s, new_gs, new_pb));
|
||||
}
|
||||
|
||||
tactic simplify_tactic(unsigned num_ns, name const * ns, options const & opts) {
|
||||
std::vector<name> names(ns, ns + num_ns);
|
||||
return mk_tactic01([=](ro_environment const & env, io_state const & ios, proof_state const & s) -> optional<proof_state> {
|
||||
return simplify_tactic(env, ios, s, names.size(), names.data(), opts);
|
||||
});
|
||||
}
|
||||
|
||||
static int mk_simplify_tactic(lua_State * L) {
|
||||
int nargs = lua_gettop(L);
|
||||
if (nargs == 0) {
|
||||
return push_tactic(L, simplify_tactic(1, &get_default_rewrite_rule_set_id(), options()));
|
||||
} else if (nargs == 1 && is_options(L, 1)) {
|
||||
return push_tactic(L, simplify_tactic(1, &get_default_rewrite_rule_set_id(), to_options(L, 1)));
|
||||
} else {
|
||||
buffer<name> rs;
|
||||
if (lua_isstring(L, 1)) {
|
||||
rs.push_back(lua_tostring(L, 1));
|
||||
} else if (lua_istable(L, 1)) {
|
||||
int n = objlen(L, 1);
|
||||
for (int i = 1; i <= n; i++) {
|
||||
lua_rawgeti(L, 1, i);
|
||||
rs.push_back(to_name_ext(L, -1));
|
||||
lua_pop(L, 1);
|
||||
}
|
||||
} else {
|
||||
rs.push_back(to_name_ext(L, 1));
|
||||
}
|
||||
options opts;
|
||||
if (nargs >= 2)
|
||||
opts = to_options(L, 2);
|
||||
return push_tactic(L, simplify_tactic(rs.size(), rs.data(), opts));
|
||||
}
|
||||
}
|
||||
|
||||
void open_simplify_tactic(lua_State * L) {
|
||||
SET_GLOBAL_FUN(mk_simplify_tactic, "simp_tac");
|
||||
}
|
||||
}
|
12
src/library/tactic/simplify_tactic.h
Normal file
12
src/library/tactic/simplify_tactic.h
Normal file
|
@ -0,0 +1,12 @@
|
|||
/*
|
||||
Copyright (c) 2013 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#pragma once
|
||||
#include "library/tactic/tactic.h"
|
||||
namespace lean {
|
||||
tactic simplify_tactic(options const & opts, unsigned num_ns, name const * ns);
|
||||
void open_simplify_tactic(lua_State * L);
|
||||
}
|
|
@ -151,6 +151,23 @@ options join(options const & opts1, options const & opts2) {
|
|||
return options(r);
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Return a new set of options based on \c opts by adding the prefix \c prefix.
|
||||
|
||||
The procedure throws an exception if \c opts contains an options (o, v), s.t. prefix + o is
|
||||
an unknow option in Lean.
|
||||
*/
|
||||
options add_prefix(name const & prefix, options const & opts) {
|
||||
std::cout << "prefix: " << prefix << "\n";
|
||||
option_declarations const & decls = get_option_declarations();
|
||||
return map(opts.m_value, [&](sexpr const & p) {
|
||||
name n = prefix + to_name(car(p));
|
||||
if (decls.find(n) == decls.end())
|
||||
throw exception(sstream() << "unknown option '" << n << "'");
|
||||
return cons(sexpr(n), cdr(p));
|
||||
});
|
||||
}
|
||||
|
||||
format pp(options const & o) {
|
||||
bool unicode = get_pp_unicode(o);
|
||||
format r;
|
||||
|
@ -182,13 +199,13 @@ std::ostream & operator<<(std::ostream & out, options const & o) {
|
|||
|
||||
DECL_UDATA(options)
|
||||
|
||||
static int mk_options(lua_State * L) {
|
||||
int mk_options(name const & prefix, lua_State * L) {
|
||||
options r;
|
||||
int nargs = lua_gettop(L);
|
||||
if (nargs % 2 != 0)
|
||||
throw exception("options expects an even number of arguments");
|
||||
for (int i = 1; i < nargs; i+=2) {
|
||||
name k = to_name_ext(L, i);
|
||||
name k = prefix + to_name_ext(L, i);
|
||||
auto it = get_option_declarations().find(k);
|
||||
if (it == get_option_declarations().end()) {
|
||||
throw exception(sstream() << "unknown option '" << k.to_string().c_str() << "'");
|
||||
|
@ -207,6 +224,10 @@ static int mk_options(lua_State * L) {
|
|||
return push_options(L, r);
|
||||
}
|
||||
|
||||
static int mk_options(lua_State * L) {
|
||||
return mk_options(name(), L);
|
||||
}
|
||||
|
||||
static int options_tostring(lua_State * L) {
|
||||
std::ostringstream out;
|
||||
out << to_options(L, 1);
|
||||
|
@ -284,6 +305,10 @@ static int options_update_string(lua_State * L) {
|
|||
return push_options(L, to_options(L, 1).update(to_name_ext(L, 2), lua_tostring(L, 3)));
|
||||
}
|
||||
|
||||
static int options_join(lua_State * L) {
|
||||
return push_options(L, join(to_options(L, 1), to_options(L, 2)));
|
||||
}
|
||||
|
||||
static int options_get(lua_State * L) {
|
||||
name k = to_name_ext(L, 2);
|
||||
auto it = get_option_declarations().find(k);
|
||||
|
@ -329,6 +354,7 @@ static const struct luaL_Reg options_m[] = {
|
|||
{"empty", safe_function<options_empty>},
|
||||
{"get", safe_function<options_get>},
|
||||
{"update", safe_function<options_update>},
|
||||
{"join", safe_function<options_join>},
|
||||
// low-level API
|
||||
{"get_bool", safe_function<options_get_bool>},
|
||||
{"get_int", safe_function<options_get_int>},
|
||||
|
|
|
@ -62,6 +62,14 @@ public:
|
|||
*/
|
||||
friend options join(options const & opts1, options const & opts2);
|
||||
|
||||
/**
|
||||
\brief Return a new set of options based on \c opts by adding the prefix \c prefix.
|
||||
|
||||
The procedure throws an exception if \c opts contains an options (o, v), s.t. prefix + o is
|
||||
an unknow option in Lean.
|
||||
*/
|
||||
friend options add_prefix(name const & prefix, options const & opts);
|
||||
|
||||
friend format pp(options const & o);
|
||||
friend std::ostream & operator<<(std::ostream & out, options const & o);
|
||||
|
||||
|
@ -80,6 +88,8 @@ struct mk_option_declaration {
|
|||
mk_option_declaration(name const & n, option_kind k, char const * default_value, char const * description);
|
||||
};
|
||||
|
||||
int mk_options(name const & prefix, lua_State * L);
|
||||
|
||||
#define LEAN_MAKE_OPTION_NAME_CORE(LINE) LEAN_OPTION_ ## LINE
|
||||
#define LEAN_MAKE_OPTION_NAME(LINE) LEAN_MAKE_OPTION_NAME_CORE(LINE)
|
||||
#define LEAN_OPTION_UNIQUE_NAME LEAN_MAKE_OPTION_NAME(__LINE__)
|
||||
|
|
8
tests/lean/simp_tac.lean
Normal file
8
tests/lean/simp_tac.lean
Normal file
|
@ -0,0 +1,8 @@
|
|||
import tactic
|
||||
|
||||
-- Create a simple rewrite set
|
||||
rewrite_set simple
|
||||
add_rewrite Nat::add_zeror Nat::add_zerol Nat::add_comm eq_id : simple
|
||||
|
||||
-- Prove the following theorem using the simplifier with the rewrite set simple
|
||||
theorem Test (a b : Nat) : 0 + a + 0 + b = b + a := (by simp simple)
|
4
tests/lean/simp_tac.lean.expected.out
Normal file
4
tests/lean/simp_tac.lean.expected.out
Normal file
|
@ -0,0 +1,4 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'tactic'
|
||||
Proved: Test
|
Loading…
Reference in a new issue