feat(frontends/lean/builtin_exprs): change how 'show' is processed in tactics

Unresolved placeholders were not being reported
This commit is contained in:
Leonardo de Moura 2015-05-19 16:21:15 -07:00
parent 1665ee39e8
commit c133d26505
4 changed files with 29 additions and 10 deletions

View file

@ -205,17 +205,16 @@ namespace sigma
intro u,
cases u with a b,
apply (sigma_eq (left_inv f a)),
show transport B (left_inv f a) ((g (f⁻¹ (f a)))⁻¹ (transport B' (right_inv f (f a))⁻¹ (g a b))) = b,
from calc
calc
transport B (left_inv f a) ((g (f⁻¹ (f a)))⁻¹ (transport B' (right_inv f (f a))⁻¹ (g a b)))
= (g a)⁻¹ (transport (B' ∘ f) (left_inv f a) (transport B' (right_inv f (f a))⁻¹ (g a b)))
: by rewrite (fn_tr_eq_tr_fn (left_inv f a) (λ a, (g a)⁻¹))
: by esimp; rewrite (fn_tr_eq_tr_fn (left_inv f a) (λ a, (g a)⁻¹))
... = (g a)⁻¹ (transport B' (ap f (left_inv f a)) (transport B' (right_inv f (f a))⁻¹ (g a b)))
: ap (g a)⁻¹ !transport_compose
... = (g a)⁻¹ (transport B' (ap f (left_inv f a)) (transport B' (ap f (left_inv f a))⁻¹ (g a b)))
: ap (λ x, (g a)⁻¹ (transport B' (ap f (left_inv f a)) (transport B' x⁻¹ (g a b)))) (adj f a)
... = (g a)⁻¹ (g a b) : {!tr_inv_tr}
... = b : by rewrite (left_inv (g a) b)
... = b : by esimp; rewrite (left_inv (g a) b)
end
definition sigma_equiv_sigma_of_is_equiv [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)]

View file

@ -228,14 +228,9 @@ static expr parse_begin_end_core(parser & p, pos_info const & pos, name const &
throw parser_error("invalid 'have' tactic, 'by', 'begin', 'proof', or 'from' expected", p.pos());
}
}
} else if (p.curr_is_token(get_show_tk())) {
auto pos = p.pos();
expr t = p.parse_tactic_expr_arg();
t = p.mk_app(get_rexact_tac_fn(), t, pos);
add_tac(t, pos);
} else if (p.curr_is_token(get_match_tk()) || p.curr_is_token(get_assume_tk()) ||
p.curr_is_token(get_take_tk()) || p.curr_is_token(get_fun_tk()) ||
p.curr_is_token(get_calc_tk())) {
p.curr_is_token(get_calc_tk()) || p.curr_is_token(get_show_tk())) {
auto pos = p.pos();
expr t = p.parse_tactic_expr_arg();
t = p.mk_app(get_exact_tac_fn(), t, pos);

7
tests/lean/show_tac.lean Normal file
View file

@ -0,0 +1,7 @@
example (a b : Prop) : a ∧ b → b ∧ a :=
begin
intro Hab,
have Ha : a, from and.elim_left Hab,
have Hb : b, from and.elim_right Hab,
show b ∧ a, from and.intro _ Ha
end

View file

@ -0,0 +1,18 @@
show_tac.lean:6:29: error: don't know how to synthesize placeholder
a b : Prop,
Hab : a ∧ b,
Ha : a,
Hb : b
⊢ b
show_tac.lean:6:2: error:invalid 'exact' tactic, term still contains metavariables after elaboration
show b ∧ a, from and.intro ?M_1 Ha
proof state:
a b : Prop,
Hab : a ∧ b,
Ha : a,
Hb : b
⊢ b ∧ a
show_tac.lean:7:0: error: failed to add declaration '14.0' to environment, value has metavariables
remark: set 'formatter.hide_full_terms' to false to see the complete term
λ (a b : Prop) (Hab : …),
?M_1