lean2/tests/lean/hott/cases_eq.hlean
Leonardo de Moura 5efadb09cc feat(library/tactic/inversion_tactic): improve 'cases' tactic for HoTT library
This commit adds support for hypotheses (h : C As idxs) where the indices idxs
are just local constants. Before this commit the indices idxs had to be hsets.
Now, they can be hsets or local constants.

The new tests demonstrate new examples that can be handled by the
improved tactic in the HoTT library
2014-12-21 15:19:25 -08:00

21 lines
508 B
Text

open eq.ops
theorem trans {A : Type} {a b c : A} (h₁ : a = b) (h₂ : b = c) : a = c :=
begin
cases h₁, cases h₂, apply rfl
end
theorem symm {A : Type} {a b : A} (h₁ : a = b) : b = a :=
begin
cases h₁, apply rfl
end
theorem congr {A B : Type} (f : A → B) {a₁ a₂ : A} (h : a₁ = a₂) : f a₁ = f a₂ :=
begin
cases h, apply rfl
end
definition inv_pV_2 {A : Type} {x y z : A} (p : x = y) (q : z = y) : (p ⬝ q⁻¹)⁻¹ = q ⬝ p⁻¹ :=
begin
cases p, cases q, apply rfl
end