feat(hit): start using induction tactic

This commit is contained in:
Floris van Doorn 2015-05-21 00:16:23 -04:00 committed by Leonardo de Moura
parent 50290fb81c
commit 40d5f83851
5 changed files with 19 additions and 44 deletions

View file

@ -25,13 +25,13 @@ namespace circle
definition rec2 {P : circle → Type} (Pb1 : P base1) (Pb2 : P base2) definition rec2 {P : circle → Type} (Pb1 : P base1) (Pb2 : P base2)
(Ps1 : seg1 ▸ Pb1 = Pb2) (Ps2 : seg2 ▸ Pb1 = Pb2) (x : circle) : P x := (Ps1 : seg1 ▸ Pb1 = Pb2) (Ps2 : seg2 ▸ Pb1 = Pb2) (x : circle) : P x :=
begin begin
fapply (suspension.rec_on x), induction x with b,
{ exact Pb1}, { exact Pb1},
{ exact Pb2}, { exact Pb2},
{ esimp, intro b, fapply (suspension.rec_on b), { esimp at *, induction b with y,
{ exact Ps1}, { exact Ps1},
{ exact Ps2}, { exact Ps2},
{ intro x, cases x}}, { cases y}},
end end
definition rec2_on [reducible] {P : circle → Type} (x : circle) (Pb1 : P base1) (Pb2 : P base2) definition rec2_on [reducible] {P : circle → Type} (x : circle) (Pb1 : P base1) (Pb2 : P base2)
@ -198,7 +198,7 @@ namespace circle
protected definition decode {x : circle} : circle.code x → base = x := protected definition decode {x : circle} : circle.code x → base = x :=
begin begin
refine circle.rec_on x _ _, induction x,
{ exact power loop}, { exact power loop},
{ apply eq_of_homotopy, intro a, { apply eq_of_homotopy, intro a,
refine !arrow.arrow_transport ⬝ !transport_eq_r ⬝ _, refine !arrow.arrow_transport ⬝ !transport_eq_r ⬝ _,
@ -208,7 +208,7 @@ namespace circle
--remove this theorem after #484 --remove this theorem after #484
theorem encode_decode {x : circle} : Π(a : circle.code x), circle.encode (circle.decode a) = a := theorem encode_decode {x : circle} : Π(a : circle.code x), circle.encode (circle.decode a) = a :=
begin begin
unfold circle.decode, refine circle.rec_on x _ _, unfold circle.decode, induction x,
{ intro a, esimp [base,base1], --simplify after #587 { intro a, esimp [base,base1], --simplify after #587
apply rec_nat_on a, apply rec_nat_on a,
{ exact idp}, { exact idp},

View file

@ -33,9 +33,9 @@ parameters {A B : Type.{u}} (f g : A → B)
protected definition rec {P : coeq → Type} (P_i : Π(x : B), P (coeq_i x)) protected definition rec {P : coeq → Type} (P_i : Π(x : B), P (coeq_i x))
(Pcp : Π(x : A), cp x ▸ P_i (f x) = P_i (g x)) (y : coeq) : P y := (Pcp : Π(x : A), cp x ▸ P_i (f x) = P_i (g x)) (y : coeq) : P y :=
begin begin
fapply (type_quotient.rec_on y), induction y,
{ intro a, apply P_i}, { apply P_i},
{ intro a a' H, cases H, apply Pcp} { cases H, apply Pcp}
end end
protected definition rec_on [reducible] {P : coeq → Type} (y : coeq) protected definition rec_on [reducible] {P : coeq → Type} (y : coeq)

View file

@ -37,11 +37,11 @@ parameters {A B : Type.{u}} (f : A → B)
(Pbase : Π(b : B), P (base b)) (Ptop : Π(a : A), P (top a)) (Pbase : Π(b : B), P (base b)) (Ptop : Π(a : A), P (top a))
(Pseg : Π(a : A), seg a ▸ Pbase (f a) = Ptop a) (x : cylinder) : P x := (Pseg : Π(a : A), seg a ▸ Pbase (f a) = Ptop a) (x : cylinder) : P x :=
begin begin
fapply (type_quotient.rec_on x), induction x,
{ intro a, cases a, { cases a,
apply Pbase, apply Pbase,
apply Ptop}, apply Ptop},
{ intro a a' H, cases H, apply Pseg} { cases H, apply Pseg}
end end
protected definition rec_on [reducible] {P : cylinder → Type} (x : cylinder) protected definition rec_on [reducible] {P : cylinder → Type} (x : cylinder)

View file

@ -36,11 +36,11 @@ parameters {TL BL TR : Type} (f : TL → BL) (g : TL → TR)
(Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▸ Pinl (f x) = Pinr (g x)) (Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▸ Pinl (f x) = Pinr (g x))
(y : pushout) : P y := (y : pushout) : P y :=
begin begin
fapply (type_quotient.rec_on y), induction y,
{ intro a, cases a, { cases a,
apply Pinl, apply Pinl,
apply Pinr}, apply Pinr},
{ intro a a' H, cases H, apply Pglue} { cases H, apply Pglue}
end end
protected definition rec_on [reducible] {P : pushout → Type} (y : pushout) protected definition rec_on [reducible] {P : pushout → Type} (y : pushout)
@ -83,31 +83,6 @@ parameters {TL BL TR : Type} (f : TL → BL) (g : TL → TR)
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_glue];apply cast_ua_fn by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_glue];apply cast_ua_fn
end end
namespace test
open pushout equiv is_equiv unit bool
private definition unit_of_empty (u : empty) : unit := star
example : pushout unit_of_empty unit_of_empty ≃ bool :=
begin
fapply equiv.MK,
{ intro x, fapply (pushout.rec_on _ _ x),
intro u, exact ff,
intro u, exact tt,
intro c, cases c},
{ intro b, cases b,
exact (inl _ _ ⋆),
exact (inr _ _ ⋆)},
{ intro b, cases b, esimp, esimp},
{ intro x, fapply (pushout.rec_on _ _ x),
intro u, cases u, esimp,
intro u, cases u, esimp,
intro c, cases c},
end
end test
end pushout end pushout
attribute pushout.inl pushout.inr [constructor] attribute pushout.inl pushout.inr [constructor]

View file

@ -72,5 +72,5 @@ end
end quotient end quotient
attribute quotient.class_of [constructor] attribute quotient.class_of [constructor]
attribute quotient.elim quotient.rec [unfold-c 7] [recursor] attribute quotient.elim quotient.rec [unfold-c 7] [recursor 7]
attribute quotient.rec_on quotient.elim_on [unfold-c 4] attribute quotient.rec_on quotient.elim_on [unfold-c 4]