lean2/tests/lean/run/tactic27.lean
Leonardo de Moura a1601e7a5f feat(library/tactic/apply_tactic): add option for 'refreshing' universe metavariables in the 'apply' tactic
The new test ../../tests/lean/run/tactic27.lean demonstrates why we need this feature. The tactic 'apply @refl' is actually 'apply @refl.{?l}'. It is used inside of a repeat tactical. Each iteration of the 'repeat' may need to use a different value for ?l. Before this commit, there was not way to say we want a fresh ?l each iteration.

Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-07-08 16:55:23 -07:00

9 lines
252 B
Text

import standard
using tactic
definition my_tac := repeat ([ apply @and_intro
| apply @refl
])
tactic_hint my_tac
theorem T1 {A : Type} {B : Type} (a : A) (b c : B) : a = a ∧ b = b ∧ c = c