fix(library/elaborator): fix glitches in the elaborator that were forcing us to provide parameters explicitly
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
1ec01f5757
commit
24528ff685
11 changed files with 36 additions and 8 deletions
Binary file not shown.
|
@ -68,7 +68,7 @@ theorem inr_inj {A B : (Type U)} {b1 b2 : B} : inr A b1 = inr A b2 → b1 = b2
|
|||
show b1 = b2,
|
||||
from optional::injectivity
|
||||
(calc some b1 = proj2 (pair (@none A) (some b1)) : refl (some b1)
|
||||
... = proj2 (pair (@none A) (some b2)) : @proj2_congr _ _ (pair (@none A) (some b1)) (pair (@none A) (some b2)) rep_eq
|
||||
... = proj2 (pair (@none A) (some b2)) : proj2_congr rep_eq
|
||||
... = some b2 : refl (some b2))
|
||||
|
||||
theorem distinct {A B : (Type U)} (a : A) (b : B) : inl a B ≠ inr A b
|
||||
|
|
|
@ -41,6 +41,7 @@ struct choice_value : public value {
|
|||
virtual expr get_type() const { lean_unreachable(); } // LCOV_EXCL_LINE
|
||||
virtual void write(serializer & ) const { lean_unreachable(); } // LCOV_EXCL_LINE
|
||||
virtual name get_name() const { return name("Choice"); }
|
||||
virtual bool is_atomic_pp(bool /* unicode */, bool /* coercion */) const { return false; } // NOLINT
|
||||
virtual void display(std::ostream & out) const {
|
||||
out << "(Choice";
|
||||
for (auto c : m_choices) {
|
||||
|
|
|
@ -262,7 +262,7 @@ class pp_fn {
|
|||
case expr_kind::Var: case expr_kind::Constant: case expr_kind::Type:
|
||||
return true;
|
||||
case expr_kind::Value:
|
||||
return !is_choice(e);
|
||||
return to_value(e).is_atomic_pp(m_unicode, m_coercion);
|
||||
case expr_kind::MetaVar:
|
||||
return !metavar_lctx(e);
|
||||
case expr_kind::App:
|
||||
|
|
|
@ -331,6 +331,7 @@ public:
|
|||
virtual void display(std::ostream & out) const;
|
||||
virtual format pp() const;
|
||||
virtual format pp(bool unicode, bool coercion) const;
|
||||
virtual bool is_atomic_pp(bool unicode, bool coercion) const = 0;
|
||||
virtual int push_lua(lua_State * L) const;
|
||||
virtual unsigned hash() const;
|
||||
virtual void write(serializer & s) const = 0;
|
||||
|
|
|
@ -67,6 +67,7 @@ public:
|
|||
expr const & get_expr() const { return m_expr; }
|
||||
context const & get_context() const { return m_ctx; }
|
||||
value_stack const & get_stack() const { return m_stack; }
|
||||
virtual bool is_atomic_pp(bool /* unicode */, bool /* coercion */) const { return false; } // NOLINT
|
||||
};
|
||||
|
||||
expr mk_closure(expr const & e, context const & ctx, value_stack const & s) { return mk_value(*(new closure(e, ctx, s))); }
|
||||
|
|
|
@ -17,6 +17,7 @@ public:
|
|||
named_value(name const & n):m_name(n) {}
|
||||
virtual ~named_value() {}
|
||||
virtual name get_name() const { return m_name; }
|
||||
virtual bool is_atomic_pp(bool /* unicode */, bool /* coercion */) const { return true; } // NOLINT
|
||||
};
|
||||
|
||||
/**
|
||||
|
|
|
@ -44,6 +44,7 @@ public:
|
|||
else
|
||||
return format(m_val);
|
||||
}
|
||||
virtual bool is_atomic_pp(bool /* unicode */, bool coercion) const { return !coercion || m_val < 0; }
|
||||
virtual unsigned hash() const { return m_val.hash(); }
|
||||
virtual int push_lua(lua_State * L) const { return push_mpz(L, m_val); }
|
||||
mpz const & get_num() const { return m_val; }
|
||||
|
|
|
@ -36,6 +36,7 @@ public:
|
|||
virtual void display(std::ostream & out) const { out << m_val; }
|
||||
virtual format pp() const { return format(m_val); }
|
||||
virtual format pp(bool, bool) const { return pp(); }
|
||||
virtual bool is_atomic_pp(bool /* unicode */, bool /* coercion */) const { return true; } // NOLINT
|
||||
virtual unsigned hash() const { return m_val.hash(); }
|
||||
virtual int push_lua(lua_State * L) const { return push_mpz(L, m_val); }
|
||||
mpz const & get_num() const { return m_val; }
|
||||
|
|
|
@ -48,6 +48,7 @@ public:
|
|||
else
|
||||
return format(m_val);
|
||||
}
|
||||
virtual bool is_atomic_pp(bool /* unicode */, bool coercion) const { return !coercion; }
|
||||
virtual unsigned hash() const { return m_val.hash(); }
|
||||
virtual int push_lua(lua_State * L) const { return push_mpq(L, m_val); }
|
||||
mpq const & get_num() const { return m_val; }
|
||||
|
|
|
@ -538,11 +538,6 @@ class elaborator::imp {
|
|||
swap(new_a, new_b);
|
||||
push_new_constraint(is_eq(c), *new_ctx, new_a, new_b, new_jst);
|
||||
return Processed;
|
||||
} else if (!has_metavar(b)) {
|
||||
// Failure, there is no way to unify
|
||||
// ?m[lift:s:n, ...] with a term that contains variables in [s, s+n]
|
||||
m_conflict = mk_failure_justification(c);
|
||||
return Failed;
|
||||
}
|
||||
}
|
||||
}
|
||||
|
@ -562,6 +557,28 @@ class elaborator::imp {
|
|||
return Continue;
|
||||
}
|
||||
|
||||
bool check_metavar_lift(unification_constraint const & c, expr const & a, expr const & b) {
|
||||
if (is_metavar(a) && has_local_context(a) && !has_metavar(b)) {
|
||||
lean_assert(!is_assigned(a));
|
||||
local_entry const & me = head(metavar_lctx(a));
|
||||
if (me.is_lift()) {
|
||||
// a is of the form ?m[lift:s:n]
|
||||
unsigned s = me.s();
|
||||
unsigned n = me.n();
|
||||
if (has_free_var(b, s, s + n, m_state.m_menv)) {
|
||||
// Failure, there is no way to unify
|
||||
// ?m[lift:s:n, ...] with a term that contains variables in [s, s+n]
|
||||
|
||||
// In older commits, this check was being done inside of process_metavar.
|
||||
// This was incorrect. This check must be performed AFTER b is normalized.
|
||||
m_conflict = mk_failure_justification(c);
|
||||
return false;
|
||||
}
|
||||
}
|
||||
}
|
||||
return true;
|
||||
}
|
||||
|
||||
justification mk_subst_justification(unification_constraint const & c, buffer<justification> const & subst_justifications) {
|
||||
if (subst_justifications.size() == 1) {
|
||||
return justification(new substitution_justification(c, subst_justifications[0]));
|
||||
|
@ -729,6 +746,8 @@ class elaborator::imp {
|
|||
expr new_a = normalize_step(ctx, a);
|
||||
expr new_b = normalize_step(ctx, b);
|
||||
if (new_a == a && new_b == b) {
|
||||
if (is_meta(a) || is_meta(b))
|
||||
break; // we do not unfold if one of the arguments is a metavar or metavar_app
|
||||
int w_a = get_unfolding_weight(a);
|
||||
int w_b = get_unfolding_weight(b);
|
||||
if (w_a >= 0 || w_b >= 0) {
|
||||
|
@ -1687,7 +1706,9 @@ class elaborator::imp {
|
|||
}
|
||||
}
|
||||
|
||||
if (!is_meta(a) && !is_meta(b) && normalize_head(a, b, c)) { return true; }
|
||||
if (!is_meta_app(a) && !is_meta_app(b) && normalize_head(a, b, c)) { return true; }
|
||||
|
||||
if (!check_metavar_lift(c, a, b) || !check_metavar_lift(c, b, a)) { return false; }
|
||||
|
||||
if (process_simple_ho_match(ctx, a, b, true, c) ||
|
||||
process_simple_ho_match(ctx, b, a, false, c))
|
||||
|
|
Loading…
Reference in a new issue