diff --git a/hott/hit/circle.hlean b/hott/hit/circle.hlean index ad5b39e18..3c798b1b5 100644 --- a/hott/hit/circle.hlean +++ b/hott/hit/circle.hlean @@ -25,13 +25,13 @@ namespace circle definition rec2 {P : circle → Type} (Pb1 : P base1) (Pb2 : P base2) (Ps1 : seg1 ▸ Pb1 = Pb2) (Ps2 : seg2 ▸ Pb1 = Pb2) (x : circle) : P x := begin - fapply (suspension.rec_on x), + induction x with b, { exact Pb1}, { exact Pb2}, - { esimp, intro b, fapply (suspension.rec_on b), + { esimp at *, induction b with y, { exact Ps1}, { exact Ps2}, - { intro x, cases x}}, + { cases y}}, end 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 := begin - refine circle.rec_on x _ _, + induction x, { exact power loop}, { apply eq_of_homotopy, intro a, refine !arrow.arrow_transport ⬝ !transport_eq_r ⬝ _, @@ -208,7 +208,7 @@ namespace circle --remove this theorem after #484 theorem encode_decode {x : circle} : Π(a : circle.code x), circle.encode (circle.decode a) = a := begin - unfold circle.decode, refine circle.rec_on x _ _, + unfold circle.decode, induction x, { intro a, esimp [base,base1], --simplify after #587 apply rec_nat_on a, { exact idp}, diff --git a/hott/hit/coeq.hlean b/hott/hit/coeq.hlean index 0a68f5dcb..08cba18dc 100644 --- a/hott/hit/coeq.hlean +++ b/hott/hit/coeq.hlean @@ -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)) (Pcp : Π(x : A), cp x ▸ P_i (f x) = P_i (g x)) (y : coeq) : P y := begin - fapply (type_quotient.rec_on y), - { intro a, apply P_i}, - { intro a a' H, cases H, apply Pcp} + induction y, + { apply P_i}, + { cases H, apply Pcp} end protected definition rec_on [reducible] {P : coeq → Type} (y : coeq) diff --git a/hott/hit/cylinder.hlean b/hott/hit/cylinder.hlean index dc3680451..f0c7eaa2e 100644 --- a/hott/hit/cylinder.hlean +++ b/hott/hit/cylinder.hlean @@ -37,11 +37,11 @@ parameters {A B : Type.{u}} (f : A → B) (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 := begin - fapply (type_quotient.rec_on x), - { intro a, cases a, - apply Pbase, - apply Ptop}, - { intro a a' H, cases H, apply Pseg} + induction x, + { cases a, + apply Pbase, + apply Ptop}, + { cases H, apply Pseg} end protected definition rec_on [reducible] {P : cylinder → Type} (x : cylinder) diff --git a/hott/hit/pushout.hlean b/hott/hit/pushout.hlean index 6f3980140..e62bcb12b 100644 --- a/hott/hit/pushout.hlean +++ b/hott/hit/pushout.hlean @@ -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)) (y : pushout) : P y := begin - fapply (type_quotient.rec_on y), - { intro a, cases a, - apply Pinl, - apply Pinr}, - { intro a a' H, cases H, apply Pglue} + induction y, + { cases a, + apply Pinl, + apply Pinr}, + { cases H, apply Pglue} end 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 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 attribute pushout.inl pushout.inr [constructor] diff --git a/hott/hit/quotient.hlean b/hott/hit/quotient.hlean index 12f952d13..511c19c25 100644 --- a/hott/hit/quotient.hlean +++ b/hott/hit/quotient.hlean @@ -72,5 +72,5 @@ end end quotient 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]