feat(hit.trunc): replace sorry by proof
This commit is contained in:
parent
de6294a4ce
commit
6ca9635d53
2 changed files with 28 additions and 30 deletions
|
@ -77,24 +77,18 @@ namespace trunc
|
||||||
|
|
||||||
section
|
section
|
||||||
open prod.ops
|
open prod.ops
|
||||||
-- definition trunc_prod_equiv : trunc n (X × Y) ≃ trunc n X × trunc n Y :=
|
definition trunc_prod_equiv : trunc n (X × Y) ≃ trunc n X × trunc n Y :=
|
||||||
-- sorry
|
begin
|
||||||
-- equiv.MK (λpp, trunc.rec_on pp (λp, (tr p.1, tr p.2)))
|
fapply equiv.MK,
|
||||||
-- (λp, trunc.rec_on p.1 (λx, trunc.rec_on p.2 (λy, tr (x,y))))
|
{exact (λpp, trunc.rec_on pp (λp, (tr p.1, tr p.2)))},
|
||||||
-- sorry --(λp, trunc.rec_on p.1 (λx, trunc.rec_on p.2 (λy, idp)))
|
{intro p, cases p with xx yy,
|
||||||
-- (λpp, trunc.rec_on pp (λp, prod.rec_on p (λx y, idp)))
|
apply (trunc.rec_on xx), intro x,
|
||||||
|
apply (trunc.rec_on yy), intro y, exact (tr (x,y))},
|
||||||
-- begin
|
{intro p, cases p with xx yy,
|
||||||
-- fapply equiv.MK,
|
apply (trunc.rec_on xx), intro x,
|
||||||
-- apply sorry, --{exact (λpp, trunc.rec_on pp (λp, (tr p.1, tr p.2)))},
|
apply (trunc.rec_on yy), intro y, apply idp},
|
||||||
-- apply sorry, /-{intro p, cases p with xx yy,
|
{intro pp, apply (trunc.rec_on pp), intro p, cases p, apply idp}
|
||||||
-- apply (trunc.rec_on xx), intro x,
|
end
|
||||||
-- apply (trunc.rec_on yy), intro y, exact (tr (x,y))},-/
|
|
||||||
-- apply sorry, /-{intro p, cases p with xx yy,
|
|
||||||
-- apply (trunc.rec_on xx), intro x,
|
|
||||||
-- apply (trunc.rec_on yy), intro y, apply idp},-/
|
|
||||||
-- apply sorry --{intro pp, apply (trunc.rec_on pp), intro p, cases p, apply idp},
|
|
||||||
-- end
|
|
||||||
end
|
end
|
||||||
|
|
||||||
/- Propositional truncation -/
|
/- Propositional truncation -/
|
||||||
|
|
|
@ -148,25 +148,29 @@ namespace trunc
|
||||||
protected definition code (n : trunc_index) (aa aa' : trunc n.+1 A) : n-Type :=
|
protected definition code (n : trunc_index) (aa aa' : trunc n.+1 A) : n-Type :=
|
||||||
trunc.rec_on aa (λa, trunc.rec_on aa' (λa', trunctype.mk' n (trunc n (a = a'))))
|
trunc.rec_on aa (λa, trunc.rec_on aa' (λa', trunctype.mk' n (trunc n (a = a'))))
|
||||||
|
|
||||||
-- protected definition encode (n : trunc_index) (aa aa' : trunc n.+1 A) : aa = aa' → code n aa aa' :=
|
protected definition encode (n : trunc_index) (aa aa' : trunc n.+1 A) : aa = aa' → code n aa aa' :=
|
||||||
-- trunc.rec_on aa (λa, trunc.rec_on aa' (λa' p, _))
|
begin
|
||||||
|
intro p, cases p, apply (trunc.rec_on aa),
|
||||||
|
intro a, esimp [code,trunc.rec_on], exact (tr idp)
|
||||||
|
end
|
||||||
|
|
||||||
|
protected definition decode (n : trunc_index) (aa aa' : trunc n.+1 A) : code n aa aa' → aa = aa' :=
|
||||||
|
begin
|
||||||
|
eapply (trunc.rec_on aa'), eapply (trunc.rec_on aa),
|
||||||
|
intro a a' x, esimp [code, trunc.rec_on] at x,
|
||||||
|
apply (trunc.rec_on x), intro p, exact (ap tr p)
|
||||||
|
end
|
||||||
|
|
||||||
definition trunc_eq_equiv (n : trunc_index) (aa aa' : trunc n.+1 A)
|
definition trunc_eq_equiv (n : trunc_index) (aa aa' : trunc n.+1 A)
|
||||||
: aa = aa' ≃ code n aa aa' :=
|
: aa = aa' ≃ code n aa aa' :=
|
||||||
begin
|
begin
|
||||||
fapply equiv.MK,
|
fapply equiv.MK,
|
||||||
{ intro p, cases p, apply (trunc.rec_on aa),
|
{ apply encode},
|
||||||
intro a, esimp [code,trunc.rec_on], exact (tr idp)},
|
{ apply decode},
|
||||||
{ eapply (trunc.rec_on aa'), eapply (trunc.rec_on aa),
|
{ eapply (trunc.rec_on aa'), eapply (trunc.rec_on aa),
|
||||||
intro a a' x, esimp [code, trunc.rec_on] at x,
|
intro a a' x, esimp [code, trunc.rec_on] at x,
|
||||||
apply (trunc.rec_on x), intro p, exact (ap tr p)},
|
apply (trunc.rec_on x), intro p, cases p, exact idp},
|
||||||
{
|
{ intro p, cases p, apply (trunc.rec_on aa), intro a, exact idp},
|
||||||
-- apply (trunc.rec_on aa'), apply (trunc.rec_on aa),
|
|
||||||
-- intro a a' x, esimp [code, trunc.rec_on] at x,
|
|
||||||
-- apply (trunc.rec_on x), intro p,
|
|
||||||
-- cases p, esimp [trunc.rec_on,eq.cases_on,compose,id], -- apply idp --?
|
|
||||||
apply sorry},
|
|
||||||
{ intro p, cases p, apply (trunc.rec_on aa), intro a, apply sorry},
|
|
||||||
end
|
end
|
||||||
|
|
||||||
definition tr_eq_tr_equiv (n : trunc_index) (a a' : A)
|
definition tr_eq_tr_equiv (n : trunc_index) (a a' : A)
|
||||||
|
|
Loading…
Reference in a new issue