parent
727a4f5a3a
commit
e3a35ba4fd
13 changed files with 215 additions and 12 deletions
|
@ -107,6 +107,8 @@ definition blast (ls : opt_identifier_list) (ds : opt_identifier_list) : tactic
|
|||
|
||||
-- with_options_tac is just a marker for the builtin 'with_options' notation
|
||||
definition with_options_tac (o : expr) (t : tactic) : tactic := builtin
|
||||
-- with_options_tac is just a marker for the builtin 'with_attributes' notation
|
||||
definition with_attributes_tac (o : expr) (n : identifier_list) (t : tactic) : tactic := builtin
|
||||
|
||||
definition cases (h : expr) (ids : opt_identifier_list) : tactic := builtin
|
||||
|
||||
|
|
|
@ -108,6 +108,8 @@ definition blast (ls : opt_identifier_list) (ds : opt_identifier_list) : tactic
|
|||
|
||||
-- with_options_tac is just a marker for the builtin 'with_options' notation
|
||||
definition with_options_tac (o : expr) (t : tactic) : tactic := builtin
|
||||
-- with_options_tac is just a marker for the builtin 'with_attributes' notation
|
||||
definition with_attributes_tac (o : expr) (n : identifier_list) (t : tactic) : tactic := builtin
|
||||
|
||||
definition simp : tactic := #tactic with_options [blast.strategy "simp"] blast
|
||||
definition simp_nohyps : tactic := #tactic with_options [blast.strategy "simp_nohyps"] blast
|
||||
|
|
|
@ -72,7 +72,7 @@
|
|||
"xrewrite" "krewrite" "blast" "simp" "esimp" "unfold" "change" "check_expr" "contradiction"
|
||||
"exfalso" "split" "existsi" "constructor" "fconstructor" "left" "right" "injection" "congruence" "reflexivity"
|
||||
"symmetry" "transitivity" "state" "induction" "induction_using" "fail" "append"
|
||||
"substvars" "now" "with_options")
|
||||
"substvars" "now" "with_options" "with_attributes" "with_attrs")
|
||||
"lean tactics")
|
||||
(defconst lean-tactics-regexp
|
||||
(eval `(rx word-start (or ,@lean-tactics) word-end)))
|
||||
|
|
|
@ -10,4 +10,4 @@ init_module.cpp elaborator_context.cpp calc_proof_elaborator.cpp
|
|||
parse_tactic_location.cpp parse_rewrite_tactic.cpp builtin_tactics.cpp
|
||||
type_util.cpp elaborator_exception.cpp local_ref_info.cpp
|
||||
obtain_expr.cpp decl_attributes.cpp nested_declaration.cpp
|
||||
parse_with_options_tactic.cpp opt_cmd.cpp prenum.cpp)
|
||||
parse_with_options_tactic.cpp opt_cmd.cpp prenum.cpp parse_with_attributes_tactic.cpp)
|
||||
|
|
|
@ -10,6 +10,7 @@ Author: Leonardo de Moura
|
|||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/parse_rewrite_tactic.h"
|
||||
#include "frontends/lean/parse_with_options_tactic.h"
|
||||
#include "frontends/lean/parse_with_attributes_tactic.h"
|
||||
|
||||
namespace lean {
|
||||
namespace notation {
|
||||
|
@ -52,6 +53,10 @@ static expr parse_with_options_tactic_expr(parser & p, unsigned, expr const *, p
|
|||
return p.save_pos(parse_with_options_tactic(p), pos);
|
||||
}
|
||||
|
||||
static expr parse_with_attributes_tactic_expr(parser & p, unsigned, expr const *, pos_info const & pos) {
|
||||
return p.save_pos(parse_with_attributes_tactic(p), pos);
|
||||
}
|
||||
|
||||
static expr parse_generalize_tactic(parser & p, unsigned, expr const *, pos_info const & pos) {
|
||||
expr e = p.parse_tactic_expr_arg();
|
||||
name id;
|
||||
|
@ -74,15 +79,17 @@ parse_table init_tactic_nud_table() {
|
|||
expr x0 = mk_var(0);
|
||||
parse_table r;
|
||||
r = r.add({transition("(", Expr), transition(")", mk_ext_action(parse_rparen))}, x0);
|
||||
r = r.add({transition("rewrite", mk_ext_action(parse_rewrite_tactic_expr))}, x0);
|
||||
r = r.add({transition("krewrite", mk_ext_action(parse_krewrite_tactic_expr))}, x0);
|
||||
r = r.add({transition("xrewrite", mk_ext_action(parse_xrewrite_tactic_expr))}, x0);
|
||||
r = r.add({transition("esimp", mk_ext_action(parse_esimp_tactic_expr))}, x0);
|
||||
r = r.add({transition("generalize", mk_ext_action(parse_generalize_tactic))}, x0);
|
||||
r = r.add({transition("unfold", mk_ext_action(parse_unfold_tactic_expr))}, x0);
|
||||
r = r.add({transition("fold", mk_ext_action(parse_fold_tactic_expr))}, x0);
|
||||
r = r.add({transition("let", mk_ext_action(parse_let_tactic))}, x0);
|
||||
r = r.add({transition("with_options", mk_ext_action(parse_with_options_tactic_expr))}, x0);
|
||||
r = r.add({transition("rewrite", mk_ext_action(parse_rewrite_tactic_expr))}, x0);
|
||||
r = r.add({transition("krewrite", mk_ext_action(parse_krewrite_tactic_expr))}, x0);
|
||||
r = r.add({transition("xrewrite", mk_ext_action(parse_xrewrite_tactic_expr))}, x0);
|
||||
r = r.add({transition("esimp", mk_ext_action(parse_esimp_tactic_expr))}, x0);
|
||||
r = r.add({transition("generalize", mk_ext_action(parse_generalize_tactic))}, x0);
|
||||
r = r.add({transition("unfold", mk_ext_action(parse_unfold_tactic_expr))}, x0);
|
||||
r = r.add({transition("fold", mk_ext_action(parse_fold_tactic_expr))}, x0);
|
||||
r = r.add({transition("let", mk_ext_action(parse_let_tactic))}, x0);
|
||||
r = r.add({transition("with_options", mk_ext_action(parse_with_options_tactic_expr))}, x0);
|
||||
r = r.add({transition("with_attributes", mk_ext_action(parse_with_attributes_tactic_expr))}, x0);
|
||||
r = r.add({transition("with_attrs", mk_ext_action(parse_with_attributes_tactic_expr))}, x0);
|
||||
return r;
|
||||
}
|
||||
|
||||
|
|
|
@ -304,4 +304,37 @@ void decl_attributes::read(deserializer & d) {
|
|||
>> m_forward >> m_elim >> m_intro >> m_intro_bang >> m_no_pattern;
|
||||
m_unfold_hint = read_list<unsigned>(d);
|
||||
}
|
||||
|
||||
bool operator==(decl_attributes const & d1, decl_attributes const & d2) {
|
||||
return
|
||||
d1.m_is_abbrev == d2.m_is_abbrev &&
|
||||
d1.m_persistent == d2.m_persistent &&
|
||||
d1.m_is_instance == d2.m_is_instance &&
|
||||
d1.m_is_trans_instance == d2.m_is_trans_instance &&
|
||||
d1.m_is_coercion == d2.m_is_coercion &&
|
||||
d1.m_is_reducible == d2.m_is_reducible &&
|
||||
d1.m_is_irreducible == d2.m_is_irreducible &&
|
||||
d1.m_is_semireducible == d2.m_is_semireducible &&
|
||||
d1.m_is_quasireducible == d2.m_is_quasireducible &&
|
||||
d1.m_is_class == d2.m_is_class &&
|
||||
d1.m_is_parsing_only == d2.m_is_parsing_only &&
|
||||
d1.m_has_multiple_instances == d2.m_has_multiple_instances &&
|
||||
d1.m_unfold_full_hint == d2.m_unfold_full_hint &&
|
||||
d1.m_constructor_hint == d2.m_constructor_hint &&
|
||||
d1.m_symm == d2.m_symm &&
|
||||
d1.m_trans == d2.m_trans &&
|
||||
d1.m_refl == d2.m_refl &&
|
||||
d1.m_subst == d2.m_subst &&
|
||||
d1.m_recursor == d2.m_recursor &&
|
||||
d1.m_simp == d1.m_simp &&
|
||||
d1.m_congr == d2.m_congr &&
|
||||
d1.m_recursor_major_pos == d2.m_recursor_major_pos &&
|
||||
d1.m_priority == d2.m_priority &&
|
||||
d1.m_forward == d2.m_forward &&
|
||||
d1.m_elim == d2.m_elim &&
|
||||
d1.m_intro == d2.m_intro &&
|
||||
d1.m_intro_bang == d2.m_intro_bang &&
|
||||
d1.m_no_pattern == d2.m_no_pattern &&
|
||||
d1.m_unfold_hint == d2.m_unfold_hint;
|
||||
}
|
||||
}
|
||||
|
|
|
@ -47,5 +47,7 @@ public:
|
|||
bool is_parsing_only() const { return m_is_parsing_only; }
|
||||
void write(serializer & s) const;
|
||||
void read(deserializer & d);
|
||||
friend bool operator==(decl_attributes const & d1, decl_attributes const & d2);
|
||||
};
|
||||
bool operator==(decl_attributes const & d1, decl_attributes const & d2);
|
||||
}
|
||||
|
|
|
@ -32,6 +32,7 @@ Author: Leonardo de Moura
|
|||
#include "frontends/lean/decl_cmds.h"
|
||||
#include "frontends/lean/nested_declaration.h"
|
||||
#include "frontends/lean/prenum.h"
|
||||
#include "frontends/lean/parse_with_attributes_tactic.h"
|
||||
|
||||
namespace lean {
|
||||
void initialize_frontend_lean_module() {
|
||||
|
@ -63,8 +64,10 @@ void initialize_frontend_lean_module() {
|
|||
initialize_obtain_expr();
|
||||
initialize_decl_cmds();
|
||||
initialize_nested_declaration();
|
||||
initialize_with_attributes_tactic();
|
||||
}
|
||||
void finalize_frontend_lean_module() {
|
||||
finalize_with_attributes_tactic();
|
||||
finalize_nested_declaration();
|
||||
finalize_decl_cmds();
|
||||
finalize_obtain_expr();
|
||||
|
|
120
src/frontends/lean/parse_with_attributes_tactic.cpp
Normal file
120
src/frontends/lean/parse_with_attributes_tactic.cpp
Normal file
|
@ -0,0 +1,120 @@
|
|||
/*
|
||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#include <string>
|
||||
#include "library/constants.h"
|
||||
#include "library/scoped_ext.h"
|
||||
#include "library/kernel_serializer.h"
|
||||
#include "library/tactic/expr_to_tactic.h"
|
||||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/decl_attributes.h"
|
||||
|
||||
namespace lean {
|
||||
static name * g_name = nullptr;
|
||||
static std::string * g_opcode = nullptr;
|
||||
|
||||
// Auxiliary macro for wrapping decl_attributes object as an expression
|
||||
class attributes_macro_cell : public macro_definition_cell {
|
||||
decl_attributes m_decl;
|
||||
public:
|
||||
attributes_macro_cell(decl_attributes const & decl):m_decl(decl) {}
|
||||
virtual name get_name() const { return *g_name; }
|
||||
virtual void write(serializer & s) const { s << *g_opcode; m_decl.write(s); }
|
||||
virtual pair<expr, constraint_seq> check_type(expr const &, extension_context &, bool) const {
|
||||
return mk_pair(mk_constant(get_tactic_expr_name()), constraint_seq());
|
||||
}
|
||||
virtual optional<expr> expand(expr const &, extension_context &) const {
|
||||
return some_expr(mk_constant(get_tactic_expr_builtin_name()));
|
||||
}
|
||||
virtual bool operator==(macro_definition_cell const & other) const {
|
||||
if (auto o = dynamic_cast<attributes_macro_cell const *>(&other))
|
||||
return m_decl == o->m_decl;
|
||||
return false;
|
||||
}
|
||||
decl_attributes const & get_attrs() const { return m_decl; }
|
||||
};
|
||||
|
||||
expr mk_attributes_expr(decl_attributes const & d) {
|
||||
macro_definition def(new attributes_macro_cell(d));
|
||||
return mk_macro(def);
|
||||
}
|
||||
|
||||
bool is_attributes_expr(expr const & e) {
|
||||
return is_macro(e) && macro_def(e).get_name() == *g_name;
|
||||
}
|
||||
|
||||
decl_attributes const & get_attributes_expr_attributes(expr const & e) {
|
||||
lean_assert(is_attributes_expr(e));
|
||||
return static_cast<attributes_macro_cell const*>(macro_def(e).raw())->get_attrs();
|
||||
}
|
||||
|
||||
static expr * g_with_attributes_tac = nullptr;
|
||||
|
||||
expr mk_with_attributes_tactic_expr(decl_attributes const & a, buffer<name> const & cs, expr const & t) {
|
||||
return mk_app(*g_with_attributes_tac, mk_attributes_expr(a), ids_to_tactic_expr(cs), t);
|
||||
}
|
||||
|
||||
tactic with_attributes(decl_attributes const & d, list<name> const & cs, tactic const & t) {
|
||||
return tactic([=](environment const & env, io_state const & ios, proof_state const & s) -> proof_state_seq {
|
||||
environment new_env = env;
|
||||
for (name const & c : cs)
|
||||
new_env = d.apply(new_env, ios, c, get_namespace(new_env));
|
||||
return t(new_env, ios, s);
|
||||
});
|
||||
}
|
||||
|
||||
expr parse_with_attributes_tactic(parser & p) {
|
||||
buffer<name> cs;
|
||||
bool persistent = false;
|
||||
bool abbrev = false;
|
||||
decl_attributes attributes(abbrev, persistent);
|
||||
name c = p.check_constant_next("invalid 'with_attributes' tactical, constant expected");
|
||||
cs.push_back(c);
|
||||
while (p.curr_is_identifier()) {
|
||||
cs.push_back(p.check_constant_next("invalid 'with_attributes' tactical, constant expected"));
|
||||
}
|
||||
attributes.parse(p);
|
||||
expr t = p.parse_tactic(get_max_prec());
|
||||
return mk_with_attributes_tactic_expr(attributes, cs, t);
|
||||
}
|
||||
|
||||
void initialize_with_attributes_tactic() {
|
||||
g_name = new name("attributes");
|
||||
g_opcode = new std::string("ATTRS");
|
||||
register_macro_deserializer(*g_opcode,
|
||||
[](deserializer & d, unsigned num, expr const *) {
|
||||
if (num > 0)
|
||||
throw corrupted_stream_exception();
|
||||
decl_attributes attrs;
|
||||
attrs.read(d);
|
||||
return mk_attributes_expr(attrs);
|
||||
});
|
||||
|
||||
name with_attributes_tac_name{"tactic", "with_attributes_tac"};
|
||||
g_with_attributes_tac = new expr(Const(with_attributes_tac_name));
|
||||
register_tac(with_attributes_tac_name,
|
||||
[=](type_checker & tc, elaborate_fn const & fn, expr const & e, pos_info_provider const * p) {
|
||||
buffer<expr> args;
|
||||
get_app_args(e, args);
|
||||
if (args.size() != 3)
|
||||
throw expr_to_tactic_exception(e, "invalid 'with_attributes' tactical, it must have two arguments");
|
||||
check_tactic_expr(args[0], "invalid 'with_attributes' tactical, invalid argument");
|
||||
expr attrs = get_tactic_expr_expr(args[0]);
|
||||
if (!is_attributes_expr(attrs))
|
||||
throw expr_to_tactic_exception(args[0], "invalid 'with_attributes' tactical, invalid argument");
|
||||
buffer<name> cs;
|
||||
get_tactic_id_list_elements(args[1], cs, "invalid 'with_attributes' tactical, invalid argument");
|
||||
tactic t = expr_to_tactic(tc, fn, args[2], p);
|
||||
return with_attributes(get_attributes_expr_attributes(attrs), to_list(cs), t);
|
||||
});
|
||||
}
|
||||
|
||||
void finalize_with_attributes_tactic() {
|
||||
delete g_name;
|
||||
delete g_opcode;
|
||||
delete g_with_attributes_tac;
|
||||
}
|
||||
}
|
15
src/frontends/lean/parse_with_attributes_tactic.h
Normal file
15
src/frontends/lean/parse_with_attributes_tactic.h
Normal file
|
@ -0,0 +1,15 @@
|
|||
/*
|
||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#pragma once
|
||||
#include "kernel/expr.h"
|
||||
|
||||
namespace lean {
|
||||
class parser;
|
||||
expr parse_with_attributes_tactic(parser & p);
|
||||
void initialize_with_attributes_tactic();
|
||||
void finalize_with_attributes_tactic();
|
||||
}
|
|
@ -96,7 +96,7 @@ void init_token_table(token_table & t) {
|
|||
{"⊢", 0}, {"⟨", g_max_prec}, {"⟩", 0}, {"^", 0}, {"↑", 0}, {"▸", 0},
|
||||
{"using", 0}, {"|", 0}, {"!", g_max_prec}, {"?", 0}, {"with", 0}, {"...", 0}, {",", 0},
|
||||
{".", 0}, {":", 0}, {"::", 0}, {"calc", 0}, {"rewrite", 0}, {"xrewrite", 0}, {"krewrite", 0},
|
||||
{"esimp", 0}, {"fold", 0}, {"unfold", 0}, {"with_options", 0},
|
||||
{"esimp", 0}, {"fold", 0}, {"unfold", 0}, {"with_options", 0}, {"with_attributes", 0}, {"with_attrs", 0},
|
||||
{"generalize", 0}, {"as", 0}, {":=", 0}, {"--", 0}, {"#", 0}, {"#tactic", 0},
|
||||
{"(*", 0}, {"/-", 0}, {"begin", g_max_prec}, {"begin+", g_max_prec}, {"abstract", g_max_prec},
|
||||
{"proof", g_max_prec}, {"qed", 0}, {"@@", g_max_prec}, {"@", g_max_prec},
|
||||
|
|
|
@ -11,6 +11,7 @@ bool.eq_ff_of_bnot_eq_tt|eq (bool.bnot ?a) bool.tt → eq ?a bool.ff
|
|||
tactic.lettac|tactic.identifier → tactic.expr → tactic
|
||||
bool.absurd_of_eq_ff_of_eq_tt|eq ?a bool.ff → eq ?a bool.tt → ?B
|
||||
bool.eq_tt_of_ne_ff|ne ?a bool.ff → eq ?a bool.tt
|
||||
tactic.with_attributes_tac|tactic.expr → tactic.identifier_list → tactic → tactic
|
||||
bool.tt_band|∀ (a : bool), eq (bool.band bool.tt a) a
|
||||
bool.cond_tt|∀ (t e : ?A), eq (bool.cond bool.tt t e) t
|
||||
bool.ff_ne_tt|eq bool.ff bool.tt → false
|
||||
|
@ -30,6 +31,7 @@ eq_ff_of_bnot_eq_tt|eq (bnot ?a) tt → eq ?a ff
|
|||
tactic.lettac|tactic.identifier → tactic.expr → tactic
|
||||
absurd_of_eq_ff_of_eq_tt|eq ?a ff → eq ?a tt → ?B
|
||||
eq_tt_of_ne_ff|ne ?a ff → eq ?a tt
|
||||
tactic.with_attributes_tac|tactic.expr → tactic.identifier_list → tactic → tactic
|
||||
cond_tt|∀ (t e : ?A), eq (cond tt t e) t
|
||||
ff_ne_tt|eq ff tt → false
|
||||
eq_ff_of_ne_tt|ne ?a tt → eq ?a ff
|
||||
|
|
17
tests/lean/run/with_attrs1.lean
Normal file
17
tests/lean/run/with_attrs1.lean
Normal file
|
@ -0,0 +1,17 @@
|
|||
import data.nat
|
||||
open nat
|
||||
|
||||
constants f : nat → nat → nat
|
||||
axiom f_ax : ∀ a b, f a b = f b a
|
||||
|
||||
example (a b : nat) : f a b = f b a :=
|
||||
begin
|
||||
with_attrs f_ax [simp] simp
|
||||
end
|
||||
|
||||
definition g (a : nat) := f a 1
|
||||
|
||||
example (a : nat) : g a = f 1 a :=
|
||||
begin
|
||||
with_attributes g [reducible] with_attributes f_ax [simp] simp
|
||||
end
|
Loading…
Reference in a new issue