lean2/tests/lean/run/dep_subst.lean
Leonardo de Moura 68688ecdf6 fix(library/tactic/subst_tactic): in the standard mode, use dependent elimination in the subst tactic (when needed)
Before this commit, the subst tactic was producing an type incorrect
result when dependent elimination was needed (see new test for an example).
2015-06-03 17:36:04 -07:00

28 lines
743 B
Text

import data.finset
open subtype setoid finset set
inductive finite_set [class] {T : Type} (xs : set T) :=
| mk : ∀ (fxs : finset T), to_set fxs = xs → finite_set xs
definition card {T : Type} (xs : set T) [fn : finite_set xs] : nat :=
begin
induction fn,
exact finset.card fxs
end
example {T : Type} (xs : set T) [fn₁ fn₂ : finite_set xs] : @card T xs fn₁ = @card T xs fn₂ :=
begin
induction fn₁ with fxs₁ h₁,
induction fn₂ with fxs₂ h₂,
subst xs,
apply sorry
end
example {T : Type} (xs : set T) [fn₁ fn₂ : finite_set xs] : @card T xs fn₁ = @card T xs fn₂ :=
begin
induction fn₁ with fxs₁ h₁,
induction fn₂ with fxs₂ h₂,
subst xs,
let aux := to_set.inj h₂,
subst aux
end