feat(hott): use the induction tactic for trunc at some places
This commit is contained in:
parent
da5f10ce63
commit
c852f7bc71
3 changed files with 9 additions and 10 deletions
|
@ -6,7 +6,7 @@ Authors: Jakob von Raumer, Floris van Doorn
|
|||
Ported from Coq HoTT
|
||||
-/
|
||||
|
||||
import .iso ..group
|
||||
import .iso algebra.group
|
||||
|
||||
open eq is_trunc iso category algebra nat unit
|
||||
|
||||
|
|
|
@ -6,7 +6,7 @@ Authors: Floris van Doorn
|
|||
Declaration of set-quotients, i.e. quotient of a mere relation which is then set-truncated.
|
||||
-/
|
||||
|
||||
import function algebra.relation types.trunc types.eq
|
||||
import function algebra.relation types.trunc types.eq hit.quotient
|
||||
|
||||
open eq is_trunc trunc quotient equiv
|
||||
|
||||
|
@ -32,9 +32,9 @@ section
|
|||
begin
|
||||
apply (@trunc.rec_on _ _ P x),
|
||||
{ intro x', apply Pt},
|
||||
{ intro y, fapply (quotient.rec_on y),
|
||||
{ exact Pc},
|
||||
{ intros, exact pathover_of_pathover_ap P tr (Pp H)}}
|
||||
{ intro y, induction y,
|
||||
{ apply Pc},
|
||||
{ exact pathover_of_pathover_ap P tr (Pp H)}}
|
||||
end
|
||||
|
||||
protected definition rec_on [reducible] {P : set_quotient → Type} (x : set_quotient)
|
||||
|
|
|
@ -186,15 +186,14 @@ namespace trunc
|
|||
|
||||
protected definition encode (n : trunc_index) (aa aa' : trunc n.+1 A) : aa = aa' → trunc.code n aa aa' :=
|
||||
begin
|
||||
intro p, induction p, apply (trunc.rec_on aa),
|
||||
intro a, esimp [trunc.code,trunc.rec_on], exact (tr idp)
|
||||
intro p, induction p, induction aa with a, esimp [trunc.code,trunc.rec_on], exact (tr idp)
|
||||
end
|
||||
|
||||
protected definition decode (n : trunc_index) (aa aa' : trunc n.+1 A) : trunc.code n aa aa' → aa = aa' :=
|
||||
begin
|
||||
eapply (trunc.rec_on aa'), eapply (trunc.rec_on aa),
|
||||
intro a a' x, esimp [trunc.code, trunc.rec_on] at x,
|
||||
apply (trunc.rec_on x), intro p, exact (ap tr p)
|
||||
induction aa' with a', induction aa with a,
|
||||
esimp [trunc.code, trunc.rec_on], intro x,
|
||||
induction x with p, exact ap tr p,
|
||||
end
|
||||
|
||||
definition trunc_eq_equiv [constructor] (n : trunc_index) (aa aa' : trunc n.+1 A)
|
||||
|
|
Loading…
Reference in a new issue