lean2/tests/lean/run/tactic26.lean
Leonardo de Moura 8743394627 refactor(kernel/inductive): replace recursor name, use '.rec' instead of '_rec'
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-09-04 15:04:57 -07:00

24 lines
731 B
Text

import logic
open tactic inhabited
namespace foo
inductive sum (A : Type) (B : Type) : Type :=
inl : A → sum A B,
inr : B → sum A B
theorem inl_inhabited {A : Type} (B : Type) (H : inhabited A) : inhabited (sum A B)
:= inhabited_destruct H (λ a, inhabited_mk (inl B a))
theorem inr_inhabited (A : Type) {B : Type} (H : inhabited B) : inhabited (sum A B)
:= inhabited_destruct H (λ b, inhabited_mk (inr A b))
definition my_tac := fixpoint (λ t, [ apply @inl_inhabited; t
| apply @inr_inhabited; t
| apply @num.num_inhabited
])
tactic_hint [inhabited] my_tac
theorem T : inhabited (sum false num.num)
end foo