diff --git a/hott/init/path.hlean b/hott/init/path.hlean index 6d85123f8..126614750 100644 --- a/hott/init/path.hlean +++ b/hott/init/path.hlean @@ -20,8 +20,8 @@ namespace eq --notation a = b := eq a b notation x = y `:>`:50 A:49 := @eq A x y - definition idp [reducible] [unfold-m] {a : A} := refl a - definition idpath [reducible] [unfold-m] (a : A) := refl a + definition idp [reducible] [constructor] {a : A} := refl a + definition idpath [reducible] [constructor] (a : A) := refl a -- unbased path induction definition rec' [reducible] {P : Π (a b : A), (a = b) -> Type} diff --git a/library/init/quot.lean b/library/init/quot.lean index e69eded61..cd166dcd2 100644 --- a/library/init/quot.lean +++ b/library/init/quot.lean @@ -117,3 +117,12 @@ namespace quot q₁ (λ a, quot.rec_on_subsingleton q₂ (λ b, f a b)) end 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] diff --git a/src/emacs/lean-syntax.el b/src/emacs/lean-syntax.el index 9f515072c..f09ae0154 100644 --- a/src/emacs/lean-syntax.el +++ b/src/emacs/lean-syntax.el @@ -119,7 +119,7 @@ (,(rx symbol-start "_" symbol-end) . 'font-lock-preprocessor-face) ;; modifiers (,(rx (or "\[persistent\]" "\[notation\]" "\[visible\]" "\[instance\]" "\[class\]" "\[parsing-only\]" - "\[coercion\]" "\[unfold-f\]" "\[unfold-m\]" "\[reducible\]" + "\[coercion\]" "\[unfold-f\]" "\[constructor\]" "\[reducible\]" "\[irreducible\]" "\[semireducible\]" "\[quasireducible\]" "\[wf\]" "\[whnf\]" "\[multiple-instances\]" "\[none\]" "\[decls\]" "\[declarations\]" "\[all-transparent\]" "\[coercions\]" "\[classes\]" diff --git a/src/frontends/lean/decl_cmds.cpp b/src/frontends/lean/decl_cmds.cpp index fd1a113d6..55a6abdd5 100644 --- a/src/frontends/lean/decl_cmds.cpp +++ b/src/frontends/lean/decl_cmds.cpp @@ -354,7 +354,7 @@ struct decl_attributes { bool m_is_parsing_only; bool m_has_multiple_instances; bool m_unfold_f_hint; - bool m_unfold_m_hint; + bool m_constructor_hint; bool m_symm; bool m_trans; bool m_refl; @@ -377,7 +377,7 @@ struct decl_attributes { m_is_parsing_only = false; m_has_multiple_instances = false; m_unfold_f_hint = false; - m_unfold_m_hint = false; + m_constructor_hint = false; m_symm = false; m_trans = false; m_refl = false; @@ -486,9 +486,9 @@ struct decl_attributes { } else if (p.curr_is_token(get_unfold_f_tk())) { p.next(); 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(); - m_unfold_m_hint = true; + m_constructor_hint = true; } else if (p.curr_is_token(get_unfold_c_tk())) { p.next(); 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); if (m_unfold_f_hint) 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) env = add_symm(env, d, m_persistent); if (m_refl) diff --git a/src/frontends/lean/token_table.cpp b/src/frontends/lean/token_table.cpp index ca764516a..2ebb677c5 100644 --- a/src/frontends/lean/token_table.cpp +++ b/src/frontends/lean/token_table.cpp @@ -109,7 +109,7 @@ void init_token_table(token_table & t) { "[none]", "[class]", "[coercion]", "[reducible]", "[irreducible]", "[semireducible]", "[quasireducible]", "[parsing-only]", "[multiple-instances]", "[symm]", "[trans]", "[refl]", "[subst]", "evaluate", "check", "eval", "[wf]", "[whnf]", "[all-transparent]", "[priority", "[unfold-f]", - "[unfold-m]", "[unfold-c", "print", + "[constructor]", "[unfold-c", "print", "end", "namespace", "section", "prelude", "help", "import", "inductive", "record", "structure", "module", "universe", "universes", "local", "precedence", "reserve", "infixl", "infixr", "infix", "postfix", "prefix", "notation", diff --git a/src/frontends/lean/tokens.cpp b/src/frontends/lean/tokens.cpp index 08eab3c6a..dfc972034 100644 --- a/src/frontends/lean/tokens.cpp +++ b/src/frontends/lean/tokens.cpp @@ -106,7 +106,7 @@ static name * g_instance = nullptr; static name * g_priority = nullptr; static name * g_unfold_c = 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_reducible = nullptr; static name * g_quasireducible = nullptr; @@ -241,7 +241,7 @@ void initialize_tokens() { g_priority = new name("[priority"); g_unfold_c = new name("[unfold-c"); 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_reducible = new name("[reducible]"); g_quasireducible = new name("[quasireducible]"); @@ -314,7 +314,7 @@ void finalize_tokens() { delete g_priority; delete g_unfold_c; delete g_unfold_f; - delete g_unfold_m; + delete g_constructor; delete g_coercion; delete g_symm; 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_unfold_c_tk() { return *g_unfold_c; } 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_symm_tk() { return *g_symm; } name const & get_trans_tk() { return *g_trans; } diff --git a/src/frontends/lean/tokens.h b/src/frontends/lean/tokens.h index 725aae400..10e31db3d 100644 --- a/src/frontends/lean/tokens.h +++ b/src/frontends/lean/tokens.h @@ -108,7 +108,7 @@ name const & get_instance_tk(); name const & get_priority_tk(); name const & get_unfold_c_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_reducible_tk(); name const & get_semireducible_tk(); diff --git a/src/library/normalize.cpp b/src/library/normalize.cpp index c85004069..0ce38bedb 100644 --- a/src/library/normalize.cpp +++ b/src/library/normalize.cpp @@ -24,7 +24,7 @@ namespace lean { - unfold_c (f a_1 ... a_i ... a_n) should be unfolded 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_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 { 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_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_add_unfold_m_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_add_constructor_entry(name const & n) { return unfold_hint_entry(unfold_hint_entry::UnfoldM, true, 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 std::string * g_key = nullptr; @@ -50,7 +50,7 @@ static std::string * g_key = nullptr; struct unfold_hint_state { name_map m_unfold_c; name_set m_unfold_f; - name_set m_unfold_m; + name_set m_constructor; }; struct unfold_hint_config { @@ -73,9 +73,9 @@ struct unfold_hint_config { break; case unfold_hint_entry::UnfoldM: if (e.m_add) - s.m_unfold_m.insert(e.m_decl_name); + s.m_constructor.insert(e.m_decl_name); else - s.m_unfold_m.erase(e.m_decl_name); + s.m_constructor.erase(e.m_decl_name); 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); } -environment add_unfold_m_hint(environment const & env, name const & n, bool persistent) { - declaration const & d = env.get(n); - if (!d.is_definition() || d.is_opaque()) - 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); +environment add_constructor_hint(environment const & env, name const & n, bool persistent) { + env.get(n); + return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_add_constructor_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); - 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) { - return unfold_hint_ext::add_entry(env, get_dummy_ios(), mk_erase_unfold_m_entry(n), 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_constructor_entry(n), persistent); } void initialize_normalize() { @@ -195,8 +193,11 @@ class normalize_fn { if (is_constructor_app(m_tc.env(), e)) return some_expr(e); expr const & f = get_app_fn(e); - if (is_constant(f) && has_unfold_m_hint(m_tc.env(), const_name(f))) { - return unfold_app(m_tc.env(), e); + if (is_constant(f) && has_constructor_hint(m_tc.env(), const_name(f))) { + if (auto r = unfold_app(m_tc.env(), e)) + return r; + else + return some_expr(e); } else { return none_expr(); } diff --git a/src/library/normalize.h b/src/library/normalize.h index fb6aa9139..6cc20117d 100644 --- a/src/library/normalize.h +++ b/src/library/normalize.h @@ -55,10 +55,10 @@ optional has_unfold_f_hint(environment const & env, name const & d); /** \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 */ -environment add_unfold_m_hint(environment const & env, name const & n, bool persistent = true); -environment erase_unfold_m_hint(environment const & env, name const & n, bool persistent = true); -/** \brief Retrieve the hint added with the procedure add_unfold_m_hint. */ -optional has_unfold_m_hint(environment const & env, name const & d); +environment add_constructor_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_constructor_hint. */ +optional has_constructor_hint(environment const & env, name const & d); void initialize_normalize(); void finalize_normalize(); diff --git a/tests/lean/587.lean b/tests/lean/587.lean new file mode 100644 index 000000000..8442145a2 --- /dev/null +++ b/tests/lean/587.lean @@ -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 diff --git a/tests/lean/587.lean.expected.out b/tests/lean/587.lean.expected.out new file mode 100644 index 000000000..a73cb1530 --- /dev/null +++ b/tests/lean/587.lean.expected.out @@ -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) diff --git a/tests/lean/hott/unfoldm.hlean b/tests/lean/hott/unfoldm.hlean index 7826a5fc7..c9a79f7a1 100644 --- a/tests/lean/hott/unfoldm.hlean +++ b/tests/lean/hott/unfoldm.hlean @@ -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