feat(library/normalize,frontends/lean): rename '[unfold-m]' hint to '[constructor]', and allow it to be attached to constants
closes #587
This commit is contained in:
parent
5fdf140096
commit
b03266be70
12 changed files with 72 additions and 38 deletions
|
@ -20,8 +20,8 @@ namespace eq
|
||||||
|
|
||||||
--notation a = b := eq a b
|
--notation a = b := eq a b
|
||||||
notation x = y `:>`:50 A:49 := @eq A x y
|
notation x = y `:>`:50 A:49 := @eq A x y
|
||||||
definition idp [reducible] [unfold-m] {a : A} := refl a
|
definition idp [reducible] [constructor] {a : A} := refl a
|
||||||
definition idpath [reducible] [unfold-m] (a : A) := refl a
|
definition idpath [reducible] [constructor] (a : A) := refl a
|
||||||
|
|
||||||
-- unbased path induction
|
-- unbased path induction
|
||||||
definition rec' [reducible] {P : Π (a b : A), (a = b) -> Type}
|
definition rec' [reducible] {P : Π (a b : A), (a = b) -> Type}
|
||||||
|
|
|
@ -117,3 +117,12 @@ namespace quot
|
||||||
q₁ (λ a, quot.rec_on_subsingleton q₂ (λ b, f a b))
|
q₁ (λ a, quot.rec_on_subsingleton q₂ (λ b, f a b))
|
||||||
end
|
end
|
||||||
end quot
|
end quot
|
||||||
|
|
||||||
|
attribute quot.mk [constructor]
|
||||||
|
attribute quot.lift_on [unfold-c 4]
|
||||||
|
attribute quot.rec [unfold-c 6]
|
||||||
|
attribute quot.rec_on [unfold-c 4]
|
||||||
|
attribute quot.rec_on_subsingleton [unfold-c 5]
|
||||||
|
attribute quot.lift₂ [unfold-c 8]
|
||||||
|
attribute quot.lift_on₂ [unfold-c 6]
|
||||||
|
attribute quot.rec_on_subsingleton₂ [unfold-c 7]
|
||||||
|
|
|
@ -119,7 +119,7 @@
|
||||||
(,(rx symbol-start "_" symbol-end) . 'font-lock-preprocessor-face)
|
(,(rx symbol-start "_" symbol-end) . 'font-lock-preprocessor-face)
|
||||||
;; modifiers
|
;; modifiers
|
||||||
(,(rx (or "\[persistent\]" "\[notation\]" "\[visible\]" "\[instance\]" "\[class\]" "\[parsing-only\]"
|
(,(rx (or "\[persistent\]" "\[notation\]" "\[visible\]" "\[instance\]" "\[class\]" "\[parsing-only\]"
|
||||||
"\[coercion\]" "\[unfold-f\]" "\[unfold-m\]" "\[reducible\]"
|
"\[coercion\]" "\[unfold-f\]" "\[constructor\]" "\[reducible\]"
|
||||||
"\[irreducible\]" "\[semireducible\]" "\[quasireducible\]" "\[wf\]"
|
"\[irreducible\]" "\[semireducible\]" "\[quasireducible\]" "\[wf\]"
|
||||||
"\[whnf\]" "\[multiple-instances\]" "\[none\]"
|
"\[whnf\]" "\[multiple-instances\]" "\[none\]"
|
||||||
"\[decls\]" "\[declarations\]" "\[all-transparent\]" "\[coercions\]" "\[classes\]"
|
"\[decls\]" "\[declarations\]" "\[all-transparent\]" "\[coercions\]" "\[classes\]"
|
||||||
|
|
|
@ -354,7 +354,7 @@ struct decl_attributes {
|
||||||
bool m_is_parsing_only;
|
bool m_is_parsing_only;
|
||||||
bool m_has_multiple_instances;
|
bool m_has_multiple_instances;
|
||||||
bool m_unfold_f_hint;
|
bool m_unfold_f_hint;
|
||||||
bool m_unfold_m_hint;
|
bool m_constructor_hint;
|
||||||
bool m_symm;
|
bool m_symm;
|
||||||
bool m_trans;
|
bool m_trans;
|
||||||
bool m_refl;
|
bool m_refl;
|
||||||
|
@ -377,7 +377,7 @@ struct decl_attributes {
|
||||||
m_is_parsing_only = false;
|
m_is_parsing_only = false;
|
||||||
m_has_multiple_instances = false;
|
m_has_multiple_instances = false;
|
||||||
m_unfold_f_hint = false;
|
m_unfold_f_hint = false;
|
||||||
m_unfold_m_hint = false;
|
m_constructor_hint = false;
|
||||||
m_symm = false;
|
m_symm = false;
|
||||||
m_trans = false;
|
m_trans = false;
|
||||||
m_refl = false;
|
m_refl = false;
|
||||||
|
@ -486,9 +486,9 @@ struct decl_attributes {
|
||||||
} else if (p.curr_is_token(get_unfold_f_tk())) {
|
} else if (p.curr_is_token(get_unfold_f_tk())) {
|
||||||
p.next();
|
p.next();
|
||||||
m_unfold_f_hint = true;
|
m_unfold_f_hint = true;
|
||||||
} else if (p.curr_is_token(get_unfold_m_tk())) {
|
} else if (p.curr_is_token(get_constructor_tk())) {
|
||||||
p.next();
|
p.next();
|
||||||
m_unfold_m_hint = true;
|
m_constructor_hint = true;
|
||||||
} else if (p.curr_is_token(get_unfold_c_tk())) {
|
} else if (p.curr_is_token(get_unfold_c_tk())) {
|
||||||
p.next();
|
p.next();
|
||||||
unsigned r = p.parse_small_nat();
|
unsigned r = p.parse_small_nat();
|
||||||
|
@ -552,9 +552,9 @@ struct decl_attributes {
|
||||||
env = add_unfold_c_hint(env, d, *m_unfold_c_hint, m_persistent);
|
env = add_unfold_c_hint(env, d, *m_unfold_c_hint, m_persistent);
|
||||||
if (m_unfold_f_hint)
|
if (m_unfold_f_hint)
|
||||||
env = add_unfold_f_hint(env, d, m_persistent);
|
env = add_unfold_f_hint(env, d, m_persistent);
|
||||||
if (m_unfold_m_hint)
|
}
|
||||||
env = add_unfold_m_hint(env, d, m_persistent);
|
if (m_constructor_hint)
|
||||||
}
|
env = add_constructor_hint(env, d, m_persistent);
|
||||||
if (m_symm)
|
if (m_symm)
|
||||||
env = add_symm(env, d, m_persistent);
|
env = add_symm(env, d, m_persistent);
|
||||||
if (m_refl)
|
if (m_refl)
|
||||||
|
|
|
@ -109,7 +109,7 @@ void init_token_table(token_table & t) {
|
||||||
"[none]", "[class]", "[coercion]", "[reducible]", "[irreducible]", "[semireducible]", "[quasireducible]",
|
"[none]", "[class]", "[coercion]", "[reducible]", "[irreducible]", "[semireducible]", "[quasireducible]",
|
||||||
"[parsing-only]", "[multiple-instances]", "[symm]", "[trans]", "[refl]", "[subst]",
|
"[parsing-only]", "[multiple-instances]", "[symm]", "[trans]", "[refl]", "[subst]",
|
||||||
"evaluate", "check", "eval", "[wf]", "[whnf]", "[all-transparent]", "[priority", "[unfold-f]",
|
"evaluate", "check", "eval", "[wf]", "[whnf]", "[all-transparent]", "[priority", "[unfold-f]",
|
||||||
"[unfold-m]", "[unfold-c", "print",
|
"[constructor]", "[unfold-c", "print",
|
||||||
"end", "namespace", "section", "prelude", "help",
|
"end", "namespace", "section", "prelude", "help",
|
||||||
"import", "inductive", "record", "structure", "module", "universe", "universes", "local",
|
"import", "inductive", "record", "structure", "module", "universe", "universes", "local",
|
||||||
"precedence", "reserve", "infixl", "infixr", "infix", "postfix", "prefix", "notation",
|
"precedence", "reserve", "infixl", "infixr", "infix", "postfix", "prefix", "notation",
|
||||||
|
|
|
@ -106,7 +106,7 @@ static name * g_instance = nullptr;
|
||||||
static name * g_priority = nullptr;
|
static name * g_priority = nullptr;
|
||||||
static name * g_unfold_c = nullptr;
|
static name * g_unfold_c = nullptr;
|
||||||
static name * g_unfold_f = nullptr;
|
static name * g_unfold_f = nullptr;
|
||||||
static name * g_unfold_m = nullptr;
|
static name * g_constructor = nullptr;
|
||||||
static name * g_coercion = nullptr;
|
static name * g_coercion = nullptr;
|
||||||
static name * g_reducible = nullptr;
|
static name * g_reducible = nullptr;
|
||||||
static name * g_quasireducible = nullptr;
|
static name * g_quasireducible = nullptr;
|
||||||
|
@ -241,7 +241,7 @@ void initialize_tokens() {
|
||||||
g_priority = new name("[priority");
|
g_priority = new name("[priority");
|
||||||
g_unfold_c = new name("[unfold-c");
|
g_unfold_c = new name("[unfold-c");
|
||||||
g_unfold_f = new name("[unfold-f]");
|
g_unfold_f = new name("[unfold-f]");
|
||||||
g_unfold_m = new name("[unfold-m]");
|
g_constructor = new name("[constructor]");
|
||||||
g_coercion = new name("[coercion]");
|
g_coercion = new name("[coercion]");
|
||||||
g_reducible = new name("[reducible]");
|
g_reducible = new name("[reducible]");
|
||||||
g_quasireducible = new name("[quasireducible]");
|
g_quasireducible = new name("[quasireducible]");
|
||||||
|
@ -314,7 +314,7 @@ void finalize_tokens() {
|
||||||
delete g_priority;
|
delete g_priority;
|
||||||
delete g_unfold_c;
|
delete g_unfold_c;
|
||||||
delete g_unfold_f;
|
delete g_unfold_f;
|
||||||
delete g_unfold_m;
|
delete g_constructor;
|
||||||
delete g_coercion;
|
delete g_coercion;
|
||||||
delete g_symm;
|
delete g_symm;
|
||||||
delete g_refl;
|
delete g_refl;
|
||||||
|
@ -514,7 +514,7 @@ name const & get_instance_tk() { return *g_instance; }
|
||||||
name const & get_priority_tk() { return *g_priority; }
|
name const & get_priority_tk() { return *g_priority; }
|
||||||
name const & get_unfold_c_tk() { return *g_unfold_c; }
|
name const & get_unfold_c_tk() { return *g_unfold_c; }
|
||||||
name const & get_unfold_f_tk() { return *g_unfold_f; }
|
name const & get_unfold_f_tk() { return *g_unfold_f; }
|
||||||
name const & get_unfold_m_tk() { return *g_unfold_m; }
|
name const & get_constructor_tk() { return *g_constructor; }
|
||||||
name const & get_coercion_tk() { return *g_coercion; }
|
name const & get_coercion_tk() { return *g_coercion; }
|
||||||
name const & get_symm_tk() { return *g_symm; }
|
name const & get_symm_tk() { return *g_symm; }
|
||||||
name const & get_trans_tk() { return *g_trans; }
|
name const & get_trans_tk() { return *g_trans; }
|
||||||
|
|
|
@ -108,7 +108,7 @@ name const & get_instance_tk();
|
||||||
name const & get_priority_tk();
|
name const & get_priority_tk();
|
||||||
name const & get_unfold_c_tk();
|
name const & get_unfold_c_tk();
|
||||||
name const & get_unfold_f_tk();
|
name const & get_unfold_f_tk();
|
||||||
name const & get_unfold_m_tk();
|
name const & get_constructor_tk();
|
||||||
name const & get_coercion_tk();
|
name const & get_coercion_tk();
|
||||||
name const & get_reducible_tk();
|
name const & get_reducible_tk();
|
||||||
name const & get_semireducible_tk();
|
name const & get_semireducible_tk();
|
||||||
|
|
|
@ -24,7 +24,7 @@ namespace lean {
|
||||||
- unfold_c (f a_1 ... a_i ... a_n) should be unfolded
|
- unfold_c (f a_1 ... a_i ... a_n) should be unfolded
|
||||||
when argument a_i is a constructor.
|
when argument a_i is a constructor.
|
||||||
- unfold_f (f a_1 ... a_i ... a_n) should be unfolded when it is fully applied.
|
- unfold_f (f a_1 ... a_i ... a_n) should be unfolded when it is fully applied.
|
||||||
- unfold_m (f ...) should be unfolded when it is the major premise of a recursor-like operator
|
- constructor (f ...) should be unfolded when it is the major premise of a recursor-like operator
|
||||||
*/
|
*/
|
||||||
struct unfold_hint_entry {
|
struct unfold_hint_entry {
|
||||||
enum kind {UnfoldC, UnfoldF, UnfoldM};
|
enum kind {UnfoldC, UnfoldF, UnfoldM};
|
||||||
|
@ -41,8 +41,8 @@ unfold_hint_entry mk_add_unfold_c_entry(name const & n, unsigned idx) { return u
|
||||||
unfold_hint_entry mk_erase_unfold_c_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldC, false, n, 0); }
|
unfold_hint_entry mk_erase_unfold_c_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldC, false, n, 0); }
|
||||||
unfold_hint_entry mk_add_unfold_f_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldF, true, n, 0); }
|
unfold_hint_entry mk_add_unfold_f_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldF, true, n, 0); }
|
||||||
unfold_hint_entry mk_erase_unfold_f_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldF, false, n, 0); }
|
unfold_hint_entry mk_erase_unfold_f_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldF, false, n, 0); }
|
||||||
unfold_hint_entry mk_add_unfold_m_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldM, true, n, 0); }
|
unfold_hint_entry mk_add_constructor_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldM, true, n, 0); }
|
||||||
unfold_hint_entry mk_erase_unfold_m_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldM, false, n, 0); }
|
unfold_hint_entry mk_erase_constructor_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldM, false, n, 0); }
|
||||||
|
|
||||||
static name * g_unfold_hint_name = nullptr;
|
static name * g_unfold_hint_name = nullptr;
|
||||||
static std::string * g_key = nullptr;
|
static std::string * g_key = nullptr;
|
||||||
|
@ -50,7 +50,7 @@ static std::string * g_key = nullptr;
|
||||||
struct unfold_hint_state {
|
struct unfold_hint_state {
|
||||||
name_map<unsigned> m_unfold_c;
|
name_map<unsigned> m_unfold_c;
|
||||||
name_set m_unfold_f;
|
name_set m_unfold_f;
|
||||||
name_set m_unfold_m;
|
name_set m_constructor;
|
||||||
};
|
};
|
||||||
|
|
||||||
struct unfold_hint_config {
|
struct unfold_hint_config {
|
||||||
|
@ -73,9 +73,9 @@ struct unfold_hint_config {
|
||||||
break;
|
break;
|
||||||
case unfold_hint_entry::UnfoldM:
|
case unfold_hint_entry::UnfoldM:
|
||||||
if (e.m_add)
|
if (e.m_add)
|
||||||
s.m_unfold_m.insert(e.m_decl_name);
|
s.m_constructor.insert(e.m_decl_name);
|
||||||
else
|
else
|
||||||
s.m_unfold_m.erase(e.m_decl_name);
|
s.m_constructor.erase(e.m_decl_name);
|
||||||
break;
|
break;
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
|
@ -138,20 +138,18 @@ environment erase_unfold_f_hint(environment const & env, name const & n, bool pe
|
||||||
return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_erase_unfold_f_entry(n), persistent);
|
return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_erase_unfold_f_entry(n), persistent);
|
||||||
}
|
}
|
||||||
|
|
||||||
environment add_unfold_m_hint(environment const & env, name const & n, bool persistent) {
|
environment add_constructor_hint(environment const & env, name const & n, bool persistent) {
|
||||||
declaration const & d = env.get(n);
|
env.get(n);
|
||||||
if (!d.is_definition() || d.is_opaque())
|
return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_add_constructor_entry(n), persistent);
|
||||||
throw exception("invalid unfold-m hint, declaration must be a non-opaque definition");
|
|
||||||
return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_add_unfold_m_entry(n), persistent);
|
|
||||||
}
|
}
|
||||||
|
|
||||||
bool has_unfold_m_hint(environment const & env, name const & d) {
|
bool has_constructor_hint(environment const & env, name const & d) {
|
||||||
unfold_hint_state const & s = unfold_hint_ext::get_state(env);
|
unfold_hint_state const & s = unfold_hint_ext::get_state(env);
|
||||||
return s.m_unfold_m.contains(d);
|
return s.m_constructor.contains(d);
|
||||||
}
|
}
|
||||||
|
|
||||||
environment erase_unfold_m_hint(environment const & env, name const & n, bool persistent) {
|
environment erase_constructor_hint(environment const & env, name const & n, bool persistent) {
|
||||||
return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_erase_unfold_m_entry(n), persistent);
|
return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_erase_constructor_entry(n), persistent);
|
||||||
}
|
}
|
||||||
|
|
||||||
void initialize_normalize() {
|
void initialize_normalize() {
|
||||||
|
@ -195,8 +193,11 @@ class normalize_fn {
|
||||||
if (is_constructor_app(m_tc.env(), e))
|
if (is_constructor_app(m_tc.env(), e))
|
||||||
return some_expr(e);
|
return some_expr(e);
|
||||||
expr const & f = get_app_fn(e);
|
expr const & f = get_app_fn(e);
|
||||||
if (is_constant(f) && has_unfold_m_hint(m_tc.env(), const_name(f))) {
|
if (is_constant(f) && has_constructor_hint(m_tc.env(), const_name(f))) {
|
||||||
return unfold_app(m_tc.env(), e);
|
if (auto r = unfold_app(m_tc.env(), e))
|
||||||
|
return r;
|
||||||
|
else
|
||||||
|
return some_expr(e);
|
||||||
} else {
|
} else {
|
||||||
return none_expr();
|
return none_expr();
|
||||||
}
|
}
|
||||||
|
|
|
@ -55,10 +55,10 @@ optional<unsigned> has_unfold_f_hint(environment const & env, name const & d);
|
||||||
|
|
||||||
/** \brief unfold-m hint instructs normalizer (and simplifier) that function application
|
/** \brief unfold-m hint instructs normalizer (and simplifier) that function application
|
||||||
(f ...) should be unfolded when it is the major premise of a constructor like operator */
|
(f ...) should be unfolded when it is the major premise of a constructor like operator */
|
||||||
environment add_unfold_m_hint(environment const & env, name const & n, bool persistent = true);
|
environment add_constructor_hint(environment const & env, name const & n, bool persistent = true);
|
||||||
environment erase_unfold_m_hint(environment const & env, name const & n, bool persistent = true);
|
environment erase_constructor_hint(environment const & env, name const & n, bool persistent = true);
|
||||||
/** \brief Retrieve the hint added with the procedure add_unfold_m_hint. */
|
/** \brief Retrieve the hint added with the procedure add_constructor_hint. */
|
||||||
optional<unsigned> has_unfold_m_hint(environment const & env, name const & d);
|
optional<unsigned> has_constructor_hint(environment const & env, name const & d);
|
||||||
|
|
||||||
void initialize_normalize();
|
void initialize_normalize();
|
||||||
void finalize_normalize();
|
void finalize_normalize();
|
||||||
|
|
15
tests/lean/587.lean
Normal file
15
tests/lean/587.lean
Normal file
|
@ -0,0 +1,15 @@
|
||||||
|
open quot setoid
|
||||||
|
|
||||||
|
variables A B : Type₁
|
||||||
|
variable s : setoid A
|
||||||
|
include s
|
||||||
|
variable f : A → B
|
||||||
|
variable c : ∀ a₁ a₂, a₁ ≈ a₂ → f a₁ = f a₂
|
||||||
|
|
||||||
|
example (a : A) (g h : B → B) : g = h → g (quot.lift_on ⟦a⟧ f c) = h (f a) :=
|
||||||
|
begin
|
||||||
|
intro gh,
|
||||||
|
esimp,
|
||||||
|
state,
|
||||||
|
rewrite gh
|
||||||
|
end
|
9
tests/lean/587.lean.expected.out
Normal file
9
tests/lean/587.lean.expected.out
Normal file
|
@ -0,0 +1,9 @@
|
||||||
|
587.lean:13:2: proof state
|
||||||
|
A B : Type₁,
|
||||||
|
s : setoid A,
|
||||||
|
f : A → B,
|
||||||
|
c : ∀ (a₁ a₂ : A), a₁ ≈ a₂ → f a₁ = f a₂,
|
||||||
|
a : A,
|
||||||
|
g h : B → B,
|
||||||
|
gh : g = h
|
||||||
|
⊢ g (f a) = h (f a)
|
|
@ -1,4 +1,4 @@
|
||||||
definition rr [unfold-m] {A : Type} {a : A} := eq.refl a
|
definition rr [constructor] {A : Type} {a : A} := eq.refl a
|
||||||
|
|
||||||
constants f g : Π {A : Type}, A → A
|
constants f g : Π {A : Type}, A → A
|
||||||
|
|
||||||
|
|
Loading…
Reference in a new issue