feat(hott): remove funext as type class, add theorems to prove equalities between functors and natural transformations

This commit is contained in:
Floris van Doorn 2015-02-23 19:54:16 -05:00
parent 9201bd7ca6
commit c091acc55b
16 changed files with 349 additions and 315 deletions

View file

@ -19,9 +19,9 @@ namespace precategory
intros, apply is_trunc_pi, intros, apply b.2,
intros, intro x, exact (a_1 (a_2 x)),
intros, exact (λ (x : a.1), x),
intros, apply funext.eq_of_homotopy, intro x, apply idp,
intros, apply funext.eq_of_homotopy, intro x, apply idp,
intros, apply funext.eq_of_homotopy, intro x, apply idp,
intros, apply eq_of_homotopy, intro x, apply idp,
intros, apply eq_of_homotopy, intro x, apply idp,
intros, apply eq_of_homotopy, intro x, apply idp,
end
end precategory

View file

@ -27,9 +27,12 @@ namespace precategory
definition id [reducible] {a : ob} : hom a a := ID a
infixr `∘` := comp
infixl `⟶`:25 := hom -- input ⟶ using \--> (this is a different arrow than \-> (→))
infixl [parsing-only] `⟶`:25 := hom -- input ⟶ using \--> (this is a different arrow than \-> (→))
namespace hom
infixl `⟶` := hom -- if you open this namespace, hom a b is printed as a ⟶ b
end hom
variables {h : hom c d} {g : hom b c} {f : hom a b} {i : hom a a}
variables {h : hom c d} {g : hom b c} {f f' : hom a b} {i : hom a a}
theorem id_compose (a : ob) : ID a ∘ ID a = ID a := !id_left
@ -44,6 +47,9 @@ namespace precategory
definition homset [reducible] (x y : ob) : hset :=
hset.mk (hom x y) _
definition is_hprop_eq_hom [instance] : is_hprop (f = f') :=
!is_trunc_eq
end precategory
structure Precategory : Type :=

View file

@ -1,14 +1,17 @@
-- Copyright (c) 2014 Floris van Doorn. All rights reserved.
-- Released under Apache 2.0 license as described in the file LICENSE.
-- Authors: Floris van Doorn, Jakob von Raumer
/-
Copyright (c) 2014 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
-- This file contains basic constructions on precategories, including common precategories
Module: algebra.precategory.constructions
Authors: Floris van Doorn, Jakob von Raumer
This file contains basic constructions on precategories, including common precategories
-/
import .nat_trans
import types.prod types.sigma types.pi
open eq prod eq eq.ops equiv is_trunc funext
open eq prod eq eq.ops equiv is_trunc
namespace precategory
namespace opposite
@ -40,7 +43,7 @@ namespace precategory
begin
apply (precategory.rec_on C), intros (hom', homH', comp', ID', assoc', id_left', id_right'),
apply (ap (λassoc'', precategory.mk hom' @homH' comp' ID' assoc'' id_left' id_right')),
repeat ( apply funext.eq_of_homotopy ; intros ),
repeat (apply eq_of_homotopy ; intros ),
apply ap,
apply (@is_hset.elim), apply !homH',
end
@ -149,22 +152,65 @@ namespace precategory
definition Precategory_hset [reducible] : Precategory :=
Precategory.mk hset precategory_hset
namespace ops
infixr `×f`:30 := product.prod_functor
infixr `ᵒᵖᶠ`:max := opposite.opposite_functor
abbreviation set := Precategory_hset
end ops
section precategory_functor
variables (C D : Precategory)
definition precategory_functor [reducible] : precategory (functor C D) :=
mk (λa b, nat_trans a b)
(λ a b, @nat_trans.to_hset C D a b)
(λ a b c g f, nat_trans.compose g f)
(λ a, nat_trans.id)
(λ a b c d h g f, !nat_trans.assoc)
(λ a b f, !nat_trans.id_left)
(λ a b f, !nat_trans.id_right)
open morphism functor nat_trans
definition precategory_functor [instance] [reducible] (C D : Precategory)
: precategory (functor C D) :=
mk (λa b, nat_trans a b)
(λ a b, @nat_trans.to_hset C D a b)
(λ a b c g f, nat_trans.compose g f)
(λ a, nat_trans.id)
(λ a b c d h g f, !nat_trans.assoc)
(λ a b f, !nat_trans.id_left)
(λ a b f, !nat_trans.id_right)
definition Precategory_functor [reducible] (C D : Precategory) : Precategory :=
Mk (precategory_functor C D)
definition Precategory_functor_rev [reducible] (C D : Precategory) : Precategory :=
Precategory_functor D C
/- we prove that if a natural transformation is pointwise an iso, then it is an iso -/
variables {C D : Precategory} {F G : C ⇒ D} (η : F ⟹ G) [iso : Π(a : C), is_iso (η a)]
include iso
definition nat_trans_inverse : G ⟹ F :=
nat_trans.mk
(λc, (η c)⁻¹)
(λc d f,
begin
apply iso.con_inv_eq_of_eq_con,
apply concat, rotate_left 1, apply assoc,
apply iso.eq_inv_con_of_con_eq,
apply inverse,
apply naturality,
end)
definition nat_trans_left_inverse : nat_trans_inverse η ∘n η = nat_trans.id :=
begin
fapply (apD011 nat_trans.mk),
apply eq_of_homotopy, intro c, apply inverse_compose,
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
apply is_hset.elim
end
definition nat_trans_right_inverse : η ∘n nat_trans_inverse η = nat_trans.id :=
begin
fapply (apD011 nat_trans.mk),
apply eq_of_homotopy, intro c, apply compose_inverse,
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
apply is_hset.elim
end
definition nat_trans_is_iso.mk : is_iso η :=
is_iso.mk (nat_trans_left_inverse η) (nat_trans_right_inverse η)
end precategory_functor
namespace ops
abbreviation set := Precategory_hset
infixr `^c`:35 := Precategory_functor_rev
infixr `×f`:30 := product.prod_functor
infixr `ᵒᵖᶠ`:(max+1) := opposite.opposite_functor
end ops
end precategory

View file

@ -4,7 +4,7 @@
import .basic types.pi
open function precategory eq prod equiv is_equiv sigma sigma.ops is_trunc
open function precategory eq prod equiv is_equiv sigma sigma.ops is_trunc funext
open pi
structure functor (C D : Precategory) : Type :=
@ -14,14 +14,99 @@ structure functor (C D : Precategory) : Type :=
(respect_comp : Π {a b c : C} (g : hom b c) (f : hom a b),
homF (g ∘ f) = homF g ∘ homF f)
infixl `⇒`:25 := functor
namespace functor
infixl `⇒`:25 := functor
variables {C D E : Precategory}
attribute obF [coercion]
attribute homF [coercion]
-- The following lemmas will later be used to prove that the type of
-- precategories forms a precategory itself
protected definition compose (G : functor D E) (F : functor C D) : functor C E :=
functor.mk
(λ x, G (F x))
(λ a b f, G (F f))
(λ a, calc
G (F (ID a)) = G (ID (F a)) : {respect_id F a}
... = ID (G (F a)) : respect_id G (F a))
(λ a b c g f, calc
G (F (g ∘ f)) = G (F g ∘ F f) : respect_comp F g f
... = G (F g) ∘ G (F f) : respect_comp G (F g) (F f))
infixr `∘f`:60 := compose
definition functor_eq_mk'' {F₁ F₂ : C → D} {H₁ : Π(a b : C), hom a b → hom (F₁ a) (F₁ b)}
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
(pF : F₁ = F₂) (pH : pF ▹ H₁ = H₂)
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
apD01111 functor.mk pF pH !is_hprop.elim !is_hprop.elim
definition functor_eq_mk' {F₁ F₂ : C → D} {H₁ : Π(a b : C), hom a b → hom (F₁ a) (F₁ b)}
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
(pF : F₁ F₂) (pH : Π(a b : C) (f : hom a b), eq_of_homotopy pF ▹ (H₁ a b f) = H₂ a b f)
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
functor_eq_mk'' id₁ id₂ comp₁ comp₂ (eq_of_homotopy pF)
(eq_of_homotopy (λc, eq_of_homotopy (λc', eq_of_homotopy (λf,
begin
apply concat, rotate_left 1, exact (pH c c' f),
apply concat, rotate_left 1,
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c c') f),
apply (apD10' f),
apply concat, rotate_left 1,
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c) c'),
apply (apD10' c'),
apply concat, rotate_left 1,
exact (pi_transport_constant (eq_of_homotopy pF) H₁ c),
apply idp
end))))
definition functor_eq_mk_constant {F : C → D} {H₁ : Π(a b : C), hom a b → hom (F a) (F b)}
{H₂ : Π(a b : C), hom a b → hom (F a) (F b)} (id₁ id₂ comp₁ comp₂)
(pH : Π(a b : C) (f : hom a b), H₁ a b f = H₂ a b f)
: functor.mk F H₁ id₁ comp₁ = functor.mk F H₂ id₂ comp₂ :=
functor_eq_mk'' id₁ id₂ comp₁ comp₂ idp
(eq_of_homotopy (λc, eq_of_homotopy (λc', eq_of_homotopy (λf, pH c c' f))))
definition functor_eq_mk {F₁ F₂ : C ⇒ D} : Π(p : obF F₁ obF F₂),
(Π(a b : C) (f : hom a b), transport (λF, hom (F a) (F b)) (eq_of_homotopy p) (F₁ f) = F₂ f)
→ F₁ = F₂ :=
functor.rec_on F₁ (λO₁ H₁ id₁ comp₁, functor.rec_on F₂ (λO₂ H₂ id₂ comp₂ p q, !functor_eq_mk' q))
-- protected definition congr
-- {C : Precategory} {D : Precategory}
-- (F : C → D)
-- (foo2 : Π ⦃a b : C⦄, hom a b → hom (F a) (F b))
-- (foo3a foo3b : Π (a : C), foo2 (ID a) = ID (F a))
-- (foo4a foo4b : Π {a b c : C} (g : @hom C C b c) (f : @hom C C a b),
-- foo2 (g ∘ f) = foo2 g ∘ foo2 f)
-- (p3 : foo3a = foo3b) (p4 : @foo4a = @foo4b)
-- : functor.mk F foo2 foo3a @foo4a = functor.mk F foo2 foo3b @foo4b
-- :=
-- begin
-- apply (eq.rec_on p3), intros,
-- apply (eq.rec_on p4), intros,
-- apply idp,
-- end
protected definition assoc {A B C D : Precategory} (H : functor C D) (G : functor B C) (F : functor A B) :
H ∘f (G ∘f F) = (H ∘f G) ∘f F :=
!functor_eq_mk_constant (λa b f, idp)
protected definition id {C : Precategory} : functor C C :=
mk (λa, a) (λ a b f, f) (λ a, idp) (λ a b c f g, idp)
protected definition ID (C : Precategory) : functor C C := id
protected definition id_left (F : functor C D) : id ∘f F = F :=
functor.rec_on F (λF1 F2 F3 F4, !functor_eq_mk_constant (λa b f, idp))
protected definition id_right (F : functor C D) : F ∘f id = F :=
functor.rec_on F (λF1 F2 F3 F4, !functor_eq_mk_constant (λa b f, idp))
set_option apply.class_instance false
-- "functor C D" is equivalent to a certain sigma type
set_option unifier.max_steps 38500
protected definition sigma_char :
@ -31,26 +116,23 @@ namespace functor
(Π {a b c : C} (g : hom b c) (f : hom a b),
homF (g ∘ f) = homF g ∘ homF f)) ≃ (functor C D) :=
begin
fapply equiv.mk,
fapply equiv.MK,
{intro S, fapply functor.mk,
exact (S.1), exact (S.2.1),
exact (pr₁ S.2.2), exact (pr₂ S.2.2)},
{fapply adjointify,
{intro F,
cases F with (d1, d2, d3, d4),
exact (sigma.mk d1 (sigma.mk d2 (pair d3 (@d4))))},
{intro F,
cases F,
apply idp},
{intro S,
cases S with (d1, S2),
cases S2 with (d2, P1),
cases P1,
apply idp}},
{intro F,
cases F with (d1, d2, d3, d4),
exact (sigma.mk d1 (sigma.mk d2 (pair d3 (@d4))))},
{intro F,
cases F,
apply idp},
{intro S,
cases S with (d1, S2),
cases S2 with (d2, P1),
cases P1,
apply idp},
end
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
protected definition strict_cat_has_functor_hset
[HD : is_hset (objects D)] : is_hset (functor C D) :=
begin
@ -67,75 +149,8 @@ namespace functor
apply is_trunc_eq, apply is_trunc_succ, apply !homH},
end
-- The following lemmas will later be used to prove that the type of
-- precategories forms a precategory itself
protected definition compose (G : functor D E) (F : functor C D) : functor C E :=
functor.mk
(λ x, G (F x))
(λ a b f, G (F f))
(λ a, calc
G (F (ID a)) = G (ID (F a)) : {respect_id F a}
... = ID (G (F a)) : respect_id G (F a))
(λ a b c g f, calc
G (F (g ∘ f)) = G (F g ∘ F f) : respect_comp F g f
... = G (F g) ∘ G (F f) : respect_comp G (F g) (F f))
infixr `∘f`:60 := compose
protected theorem congr
{C : Precategory} {D : Precategory}
(F : C → D)
(foo2 : Π ⦃a b : C⦄, hom a b → hom (F a) (F b))
(foo3a foo3b : Π (a : C), foo2 (ID a) = ID (F a))
(foo4a foo4b : Π {a b c : C} (g : @hom C C b c) (f : @hom C C a b),
foo2 (g ∘ f) = foo2 g ∘ foo2 f)
(p3 : foo3a = foo3b) (p4 : @foo4a = @foo4b)
: functor.mk F foo2 foo3a @foo4a = functor.mk F foo2 foo3b @foo4b
:=
by cases p3; cases p4; apply idp
protected theorem assoc {A B C D : Precategory} (H : functor C D) (G : functor B C) (F : functor A B) :
H ∘f (G ∘f F) = (H ∘f G) ∘f F :=
begin
cases H with (H1, H2, H3, H4),
cases G with (G1, G2, G3, G4),
cases F with (F1, F2, F3, F4),
fapply functor.congr,
{apply funext.eq_of_homotopy, intro a,
apply (@is_hset.elim), apply !homH},
{apply funext.eq_of_homotopy, intro a,
repeat (apply funext.eq_of_homotopy; intros),
apply (@is_hset.elim), apply !homH},
end
protected definition id {C : Precategory} : functor C C :=
mk (λa, a) (λ a b f, f) (λ a, idp) (λ a b c f g, idp)
protected definition ID (C : Precategory) : functor C C := id
protected theorem id_left (F : functor C D) : id ∘f F = F :=
begin
cases F with (F1, F2, F3, F4),
fapply functor.congr,
{apply funext.eq_of_homotopy, intro a,
apply (@is_hset.elim), apply !homH},
{repeat (apply funext.eq_of_homotopy; intros),
apply (@is_hset.elim), apply !homH},
end
protected theorem id_right (F : functor C D) : F ∘f id = F :=
begin
cases F with (F1, F2, F3, F4),
fapply functor.congr,
{apply funext.eq_of_homotopy, intro a,
apply (@is_hset.elim), apply !homH},
{repeat (apply funext.eq_of_homotopy; intros),
apply (@is_hset.elim), apply !homH},
end
end functor
namespace precategory
open functor
@ -152,78 +167,9 @@ namespace precategory
namespace ops
notation `PreCat`:max := Precat_of_strict_cats
abbreviation PreCat := Precat_of_strict_cats
attribute precat_of_strict_precats [instance]
end ops
end precategory
namespace functor
-- open category.ops
-- universes l m
variables {C D : Precategory}
-- check hom C D
-- variables (F : C ⟶ D) (G : D ⇒ C)
-- check G ∘ F
-- check F ∘f G
-- variables (a b : C) (f : a ⟶ b)
-- check F a
-- check F b
-- check F f
-- check G (F f)
-- print "---"
-- -- check (G ∘ F) f --error
-- check (λ(x : functor C C), x) (G ∘ F) f
-- check (G ∘f F) f
-- print "---"
-- -- check (G ∘ F) a --error
-- check (G ∘f F) a
-- print "---"
-- -- check λ(H : hom C D) (x : C), H x --error
-- check λ(H : @hom _ Cat C D) (x : C), H x
-- check λ(H : C ⇒ D) (x : C), H x
-- print "---"
-- -- variables {obF obG : C → D} (Hob : ∀x, obF x = obG x) (homF : Π(a b : C) (f : a ⟶ b), obF a ⟶ obF b) (homG : Π(a b : C) (f : a ⟶ b), obG a ⟶ obG b)
-- -- check eq.rec_on (funext Hob) homF = homG
/-theorem mk_heq {obF obG : C → D} {homF homG idF idG compF compG} (Hob : ∀x, obF x = obG x)
(Hmor : ∀(a b : C) (f : a ⟶ b), homF a b f == homG a b f)
: mk obF homF idF compF = mk obG homG idG compG :=
hddcongr_arg4 mk
(funext Hob)
(hfunext (λ a, hfunext (λ b, hfunext (λ f, !Hmor))))
!proof_irrel
!proof_irrel
protected theorem hequal {F G : C ⇒ D} : Π (Hob : ∀x, F x = G x)
(Hmor : ∀a b (f : a ⟶ b), F f == G f), F = G :=
functor.rec
(λ obF homF idF compF,
functor.rec
(λ obG homG idG compG Hob Hmor, mk_heq Hob Hmor)
G)
F-/
-- theorem mk_eq {obF obG : C → D} {homF homG idF idG compF compG} (Hob : ∀x, obF x = obG x)
-- (Hmor : ∀(a b : C) (f : a ⟶ b), cast (congr_arg (λ x, x a ⟶ x b) (funext Hob)) (homF a b f)
-- = homG a b f)
-- : mk obF homF idF compF = mk obG homG idG compG :=
-- dcongr_arg4 mk
-- (funext Hob)
-- (funext (λ a, funext (λ b, funext (λ f, sorry ⬝ Hmor a b f))))
-- -- to fill this sorry use (a generalization of) cast_pull
-- !proof_irrel
-- !proof_irrel
-- protected theorem equal {F G : C ⇒ D} : Π (Hob : ∀x, F x = G x)
-- (Hmor : ∀a b (f : a ⟶ b), cast (congr_arg (λ x, x a ⟶ x b) (funext Hob)) (F f) = G f), F = G :=
-- functor.rec
-- (λ obF homF idF compF,
-- functor.rec
-- (λ obG homG idG compG Hob Hmor, mk_eq Hob Hmor)
-- G)
-- F
end functor

View file

@ -26,6 +26,8 @@ namespace morphism
is_iso.rec (λg h1 h2, g) H
postfix `⁻¹` := inverse
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ʰ`:100 := inverse
theorem inverse_compose (f : a ⟶ b) [H : is_iso f] : f⁻¹ ∘ f = id :=
is_iso.rec (λg h1 h2, h1) H

View file

@ -2,18 +2,18 @@
-- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Floris van Doorn, Jakob von Raumer
import .functor
open eq precategory functor is_trunc equiv sigma.ops sigma is_equiv function pi
import .functor .morphism
open eq precategory functor is_trunc equiv sigma.ops sigma is_equiv function pi funext
inductive nat_trans {C D : Precategory} (F G : C ⇒ D) : Type :=
mk : Π (η : Π (a : C), hom (F a) (G a))
(nat : Π {a b : C} (f : hom a b), G f ∘ η a = η b ∘ F f),
nat_trans F G
infixl `⟹`:25 := nat_trans -- \==>
namespace nat_trans
variables {C D : Precategory} {F G H I : functor C D}
infixl `⟹`:25 := nat_trans -- \==>
variables {C D : Precategory} {F G H I : C ⇒ D}
definition natural_map [coercion] (η : F ⟹ G) : Π (a : C), F a ⟶ G a :=
nat_trans.rec (λ x y, x) η
@ -34,58 +34,33 @@ namespace nat_trans
infixr `∘n`:60 := compose
protected theorem congr
{C : Precategory} {D : Precategory}
(F G : C ⇒ D)
(η₁ η₂ : Π (a : C), hom (F a) (G a))
local attribute is_hprop_eq_hom [instance]
definition nat_trans_eq_mk' {η₁ η₂ : Π (a : C), hom (F a) (G a)}
(nat₁ : Π (a b : C) (f : hom a b), G f ∘ η₁ a = η₁ b ∘ F f)
(nat₂ : Π (a b : C) (f : hom a b), G f ∘ η₂ a = η₂ b ∘ F f)
(p₁ : η₁ = η₂) (p₂ : p₁ ▹ nat₁ = nat₂)
: @nat_trans.mk C D F G η₁ nat₁ = @nat_trans.mk C D F G η₂ nat₂
:=
begin
apply (apD011 (@nat_trans.mk C D F G) p₁ p₂),
end
(p : η₁ η₂)
: nat_trans.mk η₁ nat₁ = nat_trans.mk η₂ nat₂ :=
apD011 nat_trans.mk (eq_of_homotopy p) !is_hprop.elim
set_option apply.class_instance false -- disable class instance resolution in the apply tactic
definition nat_trans_eq_mk {η₁ η₂ : F ⟹ G} : natural_map η₁ natural_map η₂ → η₁ = η₂ :=
nat_trans.rec_on η₁ (λf₁ nat₁, nat_trans.rec_on η₂ (λf₂ nat₂ p, !nat_trans_eq_mk' p))
protected definition assoc (η₃ : H ⟹ I) (η₂ : G ⟹ H) (η₁ : F ⟹ G) :
η₃ ∘n (η₂ ∘n η₁) = (η₃ ∘n η₂) ∘n η₁ :=
begin
cases η₃, cases η₂, cases η₁,
fapply nat_trans.congr,
{apply funext.eq_of_homotopy, intro a,
apply assoc},
{repeat (apply funext.eq_of_homotopy; intros),
apply (@is_hset.elim), apply !homH},
end
nat_trans_eq_mk (λa, !assoc)
protected definition id {C D : Precategory} {F : functor C D} : nat_trans F F :=
mk (λa, id) (λa b f, !id_right ⬝ (!id_left⁻¹))
mk (λa, id) (λa b f, !id_right ⬝ !id_left⁻¹)
protected definition ID {C D : Precategory} (F : functor C D) : nat_trans F F :=
id
protected definition id_left (η : F ⟹ G) : id ∘n η = η :=
begin
cases η,
fapply (nat_trans.congr F G),
{apply funext.eq_of_homotopy, intro a,
apply id_left},
{repeat (apply funext.eq_of_homotopy; intros),
apply (@is_hset.elim), apply !homH},
end
nat_trans_eq_mk (λa, !id_left)
protected definition id_right (η : F ⟹ G) : η ∘n id = η :=
begin
cases η,
fapply (nat_trans.congr F G),
{apply funext.eq_of_homotopy, intros, apply id_right},
{repeat (apply funext.eq_of_homotopy; intros),
apply (@is_hset.elim), apply !homH},
end
nat_trans_eq_mk (λa, !id_right)
--set_option pp.implicit true
protected definition sigma_char (F G : C ⇒ D) :
(Σ (η : Π (a : C), hom (F a) (G a)), Π (a b : C) (f : hom a b), G f ∘ η a = η b ∘ F f) ≃ (F ⟹ G) :=
begin
@ -96,20 +71,15 @@ namespace nat_trans
fapply sigma.mk,
intro a, exact (H a),
intros (a, b, f), exact (naturality H f),
intro H, apply (nat_trans.rec_on H),
intros (eta, nat), unfold function.id,
fapply nat_trans.congr,
apply idp,
repeat ( apply funext.eq_of_homotopy ; intros ),
apply (@is_hset.elim), apply !homH,
intro η, apply nat_trans_eq_mk, intro a, apply idp,
intro S,
fapply sigma_eq,
apply funext.eq_of_homotopy, intro a,
apply eq_of_homotopy, intro a,
apply idp,
repeat ( apply funext.eq_of_homotopy ; intros ),
apply (@is_hset.elim), apply !homH,
apply is_hprop.elim,
end
set_option apply.class_instance false
protected definition to_hset : is_hset (F ⟹ G) :=
begin
apply is_trunc_is_equiv_closed, apply (equiv.to_is_equiv !sigma_char),

View file

@ -9,7 +9,7 @@ Author: Floris van Doorn
--note: modify definition in category.set
import .constructions .morphism
open eq precategory equiv is_equiv is_trunc
open eq precategory functor is_trunc equiv is_equiv pi
open is_trunc.trunctype funext precategory.ops prod.ops
set_option pp.beta true
@ -36,42 +36,63 @@ namespace yoneda
end
end yoneda
attribute precategory_functor [instance] [reducible]
namespace nat_trans
open morphism functor
variables {C D : Precategory} {F G : C ⇒ D} (η : F ⟹ G) (H : Π(a : C), is_iso (η a))
include H
definition nat_trans_inverse : G ⟹ F :=
nat_trans.mk
(λc, (η c)⁻¹)
(λc d f,
begin
apply iso.con_inv_eq_of_eq_con,
apply concat, rotate_left 1, apply assoc,
apply iso.eq_inv_con_of_con_eq,
apply inverse,
apply naturality,
end)
definition nat_trans_left_inverse : nat_trans_inverse η H ∘ η = nat_trans.id :=
begin
fapply (apD011 nat_trans.mk),
apply eq_of_homotopy, intro c, apply inverse_compose,
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, fapply is_hset.elim
end
definition nat_trans_right_inverse : η ∘ nat_trans_inverse η H = nat_trans.id :=
begin
fapply (apD011 nat_trans.mk),
apply eq_of_homotopy, intro c, apply compose_inverse,
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, fapply is_hset.elim
end
definition nat_trans_is_iso.mk : is_iso η :=
is_iso.mk (nat_trans_left_inverse η H) (nat_trans_right_inverse η H)
namespace functor
open prod nat_trans
variables {C D E : Precategory}
end nat_trans
definition functor_curry_ob (F : C ×c D ⇒ E) (c : C) : E ^c D :=
functor.mk (λd, F (c,d))
(λd d' g, F (id, g))
(λd, !respect_id)
(λd₁ d₂ d₃ g' g, proof calc
F (id, g' ∘ g) = F (id ∘ id, g' ∘ g) : {(id_compose c)⁻¹}
... = F ((id,g') ∘ (id, g)) : idp
... = F (id,g') ∘ F (id, g) : respect_comp F qed)
local abbreviation Fob := @functor_curry_ob
definition functor_curry_mor (F : C ×c D ⇒ E) ⦃c c' : C⦄ (f : c ⟶ c') : Fob F c ⟹ Fob F c' :=
nat_trans.mk (λd, F (f, id))
(λd d' g, proof calc
F (id, g) ∘ F (f, id) = F (id ∘ f, g ∘ id) : respect_comp F
... = F (f, g ∘ id) : {id_left f}
... = F (f, g) : {id_right g}
... = F (f ∘ id, g) : {(id_right f)⁻¹}
... = F (f ∘ id, id ∘ g) : {(id_left g)⁻¹}
... = F (f, id) ∘ F (id, g) : (respect_comp F (f, id) (id, g))⁻¹ᵖ
qed)
local abbreviation Fmor := @functor_curry_mor
definition functor_curry_mor_def (F : C ×c D ⇒ E) ⦃c c' : C⦄ (f : c ⟶ c') (d : D) :
(Fmor F f) d = F (f, id) := idp
definition functor_curry_id (F : C ×c D ⇒ E) (c : C) : Fmor F (ID c) = nat_trans.id :=
nat_trans_eq_mk (λd, respect_id F _)
definition functor_curry_comp (F : C ×c D ⇒ E) ⦃c c' c'' : C⦄ (f' : c' ⟶ c'') (f : c ⟶ c')
: Fmor F (f' ∘ f) = Fmor F f' ∘n Fmor F f :=
nat_trans_eq_mk (λd, calc
natural_map (Fmor F (f' ∘ f)) d = F (f' ∘ f, id) : functor_curry_mor_def
... = F (f' ∘ f, id ∘ id) : {(id_compose d)⁻¹}
... = F ((f',id) ∘ (f, id)) : idp
... = F (f',id) ∘ F (f, id) : respect_comp F
... = natural_map ((Fmor F f') ∘ (Fmor F f)) d : idp)
--respect_comp F (g, id) (f, id)
definition functor_curry (F : C ×c D ⇒ E) : C ⇒ E ^c D :=
functor.mk (functor_curry_ob F)
(functor_curry_mor F)
(functor_curry_id F)
(functor_curry_comp F)
definition is_equiv_functor_curry : is_equiv (@functor_curry C D E) := sorry
definition equiv_functor_curry : (C ×c D ⇒ E) ≃ (C ⇒ E ^c D) :=
equiv.mk _ !is_equiv_functor_curry
end functor
-- Coq uses unit/counit definitions as basic
-- open yoneda precategory.product precategory.opposite functor morphism

View file

@ -10,9 +10,12 @@ open eq
-- ------
-- Define function extensionality as a type class
structure funext [class] : Type :=
(elim : Π (A : Type) (P : A → Type ) (f g : Π x, P x), is_equiv (@apD10 A P f g))
-- structure funext [class] : Type :=
-- (elim : Π (A : Type) (P : A → Type ) (f g : Π x, P x), is_equiv (@apD10 A P f g))
-- set_option pp.universes true
-- check @funext.mk
-- check @funext.elim
exit
namespace funext

View file

@ -126,5 +126,25 @@ theorem weak_funext_of_ua : weak_funext :=
)
-- In the following we will proof function extensionality using the univalence axiom
definition funext_of_ua [instance] : funext :=
definition funext_of_ua : funext :=
funext_of_weak_funext (@weak_funext_of_ua)
namespace funext
definition is_equiv_apD [instance] {A : Type} {P : A → Type} (f g : Π x, P x)
: is_equiv (@apD10 A P f g) :=
funext_of_ua f g
end funext
open funext
definition eq_of_homotopy {A : Type} {P : A → Type} {f g : Π x, P x} : f g → f = g :=
is_equiv.inv (@apD10 A P f g)
definition eq_of_homotopy2 {A B : Type} {P : A → B → Type}
(f g : Πx y, P x y) : (Πx y, f x y = g x y) → f = g :=
λ E, eq_of_homotopy (λx, eq_of_homotopy (E x))
definition naive_funext_of_ua : naive_funext :=
λ A P f g h, eq_of_homotopy h
exit

View file

@ -7,7 +7,7 @@ import ..path ..trunc ..equiv .funext
open eq is_trunc sigma function
/- In hott.axioms.funext, we defined function extensionality to be the assertion
/- In init.axioms.funext, we defined function extensionality to be the assertion
that the map apD10 is an equivalence. We now prove that this follows
from a couple of weaker-looking forms of function extensionality. We do
require eta conversion, which Coq 8.4+ has judgmentally.
@ -15,19 +15,17 @@ open eq is_trunc sigma function
This proof is originally due to Voevodsky; it has since been simplified
by Peter Lumsdaine and Michael Shulman. -/
definition funext.{l k} :=
Π ⦃A : Type.{l}⦄ {P : A → Type.{k}} (f g : Π x, P x), is_equiv (@apD10 A P f g)
-- Naive funext is the simple assertion that pointwise equal functions are equal.
-- TODO think about universe levels
definition naive_funext :=
Π {A : Type} {P : A → Type} (f g : Πx, P x), (f g) → f = g
Π ⦃A : Type⦄ {P : A → Type} (f g : Πx, P x), (f g) → f = g
-- Weak funext says that a product of contractible types is contractible.
definition weak_funext :=
Π {A : Type} (P : A → Type) [H: Πx, is_contr (P x)], is_contr (Πx, P x)
-- The obvious implications are Funext -> NaiveFunext -> WeakFunext
-- TODO: Get class inference to work locally
definition naive_funext_from_funext [F : funext] : naive_funext :=
(λ A P f g h, funext.eq_of_homotopy h)
Π ⦃A : Type⦄ (P : A → Type) [H: Πx, is_contr (P x)], is_contr (Πx, P x)
definition weak_funext_of_naive_funext : naive_funext → weak_funext :=
(λ nf A P (Pc : Πx, is_contr (P x)),
@ -48,7 +46,7 @@ definition weak_funext_of_naive_funext : naive_funext → weak_funext :=
context
universes l k
parameters [wf : weak_funext.{l k}] {A : Type.{l}} {B : A → Type.{k}} (f : Π x, B x)
parameters (wf : weak_funext.{l k}) {A : Type.{l}} {B : A → Type.{k}} (f : Π x, B x)
definition is_contr_sigma_homotopy [instance] : is_contr (Σ (g : Π x, B x), f g) :=
is_contr.mk (sigma.mk f (homotopy.refl f))
@ -90,7 +88,7 @@ universe variables l k
local attribute weak_funext [reducible]
theorem funext_of_weak_funext (wf : weak_funext.{l k}) : funext.{l k} :=
funext.mk (λ A B f g,
λ A B f g,
let eq_to_f := (λ g' x, f = g') in
let sim2path := homotopy_ind f eq_to_f idp in
have t1 : sim2path f (homotopy.refl f) = idp,
@ -101,7 +99,7 @@ theorem funext_of_weak_funext (wf : weak_funext.{l k}) : funext.{l k} :=
proof (homotopy_ind f (λ g' x, apD10 (sim2path g' x) = x) t2) g qed,
have retr : (sim2path g) ∘ apD10 id,
from (λ h, eq.rec_on h (homotopy_ind_comp f _ idp)),
is_equiv.adjointify apD10 (sim2path g) sect retr)
is_equiv.adjointify apD10 (sim2path g) sect retr
definition funext_from_naive_funext : naive_funext -> funext :=
compose funext_of_weak_funext weak_funext_of_naive_funext

View file

@ -32,6 +32,8 @@ structure equiv (A B : Type) :=
namespace is_equiv
/- Some instances and closure properties of equivalences -/
postfix `⁻¹` := inv
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ᵉ`:100 := inv
section
variables {A B C : Type} (f : A → B) (g : B → C) {f' : A → B}

View file

@ -42,6 +42,9 @@ namespace eq
notation p₁ ⬝ p₂ := concat p₁ p₂
notation p ⁻¹ := inverse p
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ᵖ`:100 := inverse
/- The 1-dimensional groupoid structure -/
@ -208,6 +211,10 @@ namespace eq
definition apD10 {f g : Πx, P x} (H : f = g) : f g :=
λx, eq.rec_on H idp
--the next theorem is useful if you want to write "apply (apD10' a)"
definition apD10' {f g : Πx, P x} (a : A) (H : f = g) : f a = g a :=
eq.rec_on H idp
definition ap10 {f g : A → B} (H : f = g) : f g := apD10 H
definition ap11 {f g : A → B} (H : f = g) {x y : A} (p : x = y) : f x = g y :=
@ -639,31 +646,40 @@ namespace eq
end eq
namespace eq
variables {A B C D E : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D}
section
variables {A B C D E : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D}
theorem ap011 (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' :=
eq.rec_on Ha (eq.rec_on Hb idp)
definition ap011 (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' :=
eq.rec_on Ha (eq.rec_on Hb idp)
theorem ap0111 (f : A → B → C → D) (Ha : a = a') (Hb : b = b') (Hc : c = c')
: f a b c = f a' b' c' :=
eq.rec_on Ha (ap011 (f a) Hb Hc)
definition ap0111 (f : A → B → C → D) (Ha : a = a') (Hb : b = b') (Hc : c = c')
: f a b c = f a' b' c' :=
eq.rec_on Ha (ap011 (f a) Hb Hc)
theorem ap01111 (f : A → B → C → D → E) (Ha : a = a') (Hb : b = b') (Hc : c = c') (Hd : d = d')
: f a b c d = f a' b' c' d' :=
eq.rec_on Ha (ap0111 (f a) Hb Hc Hd)
end eq
namespace eq
variables {A : Type} {B : A → Type} {C : Πa, B a → Type} {D : Πa b, C a b → Type}
{E : Πa b c, D a b c → Type} {F : Type}
variables {a a' : A}
{b : B a} {b' : B a'}
{c : C a b} {c' : C a' b'}
{d : D a b c} {d' : D a' b' c'}
theorem apD011 (f : Πa, B a → F) (Ha : a = a') (Hb : (Ha ▹ b) = b')
: f a b = f a' b' :=
eq.rec_on Hb (eq.rec_on Ha idp)
definition ap01111 (f : A → B → C → D → E) (Ha : a = a') (Hb : b = b') (Hc : c = c') (Hd : d = d')
: f a b c d = f a' b' c' d' :=
eq.rec_on Ha (ap0111 (f a) Hb Hc Hd)
end
section
variables {A : Type} {B : A → Type} {C : Πa, B a → Type} {D : Πa b, C a b → Type}
{E : Πa b c, D a b c → Type} {F : Type}
variables {a a' : A}
{b : B a} {b' : B a'}
{c : C a b} {c' : C a' b'}
{d : D a b c} {d' : D a' b' c'}
definition apD011 (f : Πa, B a → F) (Ha : a = a') (Hb : (Ha ▹ b) = b')
: f a b = f a' b' :=
eq.rec_on Hb (eq.rec_on Ha idp)
definition apD0111 (f : Πa b, C a b → F) (Ha : a = a') (Hb : (Ha ▹ b) = b')
(Hc : apD011 C Ha Hb ▹ c = c')
: f a b c = f a' b' c' :=
eq.rec_on Hc (eq.rec_on Hb (eq.rec_on Ha idp))
definition apD01111 (f : Πa b c, D a b c → F) (Ha : a = a') (Hb : (Ha ▹ b) = b')
(Hc : apD011 C Ha Hb ▹ c = c') (Hd : apD0111 D Ha Hb Hc ▹ d = d')
: f a b c d = f a' b' c' d' :=
eq.rec_on Hd (eq.rec_on Hc (eq.rec_on Hb (eq.rec_on Ha idp)))
end
end eq

View file

@ -29,7 +29,7 @@ namespace is_trunc
postfix `.+1`:(max+1) := trunc_index.succ
postfix `.+2`:(max+1) := λn, (n .+1 .+1)
notation `-2` := trunc_index.minus_two
notation `-1` := -2.+1
notation `-1` := -2.+1 -- ISSUE: -1 gets printed as -2.+1
export [coercions] nat
namespace trunc_index

View file

@ -11,8 +11,7 @@ import types.sigma
open eq equiv is_equiv funext
namespace pi
universe variables l k
variables {A A' : Type.{l}} {B : A → Type.{k}} {B' : A' → Type.{k}} {C : Πa, B a → Type}
variables {A A' : Type} {B : A → Type} {B' : A' → Type} {C : Πa, B a → Type}
{D : Πa b, C a b → Type}
{a a' a'' : A} {b b₁ b₂ : B a} {b' : B a'} {b'' : B a''} {f g : Πa, B a}
@ -36,10 +35,10 @@ namespace pi
/- The identification of the path space of a dependent function space, up to equivalence, is of course just funext. -/
definition eq_equiv_homotopy (f g : Πx, B x) : (f = g) ≃ (f g) :=
equiv.mk _ !funext.elim
equiv.mk _ !is_equiv_apD
definition is_equiv_eq_of_homotopy [instance] (f g : Πx, B x)
: is_equiv (@eq_of_homotopy _ _ _ f g) :=
: is_equiv (@eq_of_homotopy _ _ f g) :=
is_equiv_inv apD10
definition homotopy_equiv_eq (f g : Πx, B x) : (f g) ≃ (f = g) :=
@ -56,7 +55,7 @@ namespace pi
/- A special case of [transport_pi] where the type [B] does not depend on [A],
and so it is just a fixed type [B]. -/
definition pi_transport_constant {C : A → A' → Type} (p : a = a') (f : Π(b : A'), C a b)
: (transport (λa, Π(b : A'), C a b) p f) (λb, transport (λa, C a b) p (f b)) :=
: Π(b : A'), (transport (λa, Π(b : A'), C a b) p f) b = transport (λa, C a b) p (f b) :=
eq.rec_on p (λx, idp)
/- Maps on paths -/
@ -79,7 +78,7 @@ namespace pi
(g : Π(b' : B a'), C a' b') : (Π(b : B a), p ▹D (f b) = g (p ▹ b)) ≃ (p ▹ f = g) :=
eq.rec_on p (λg, !homotopy_equiv_eq) g
definition heq_pi {C : A → Type.{k}} (p : a = a') (f : Π(b : B a), C a)
definition heq_pi {C : A → Type} (p : a = a') (f : Π(b : B a), C a)
(g : Π(b' : B a'), C a') : (Π(b : B a), p ▹ (f b) = g (p ▹ b)) ≃ (p ▹ f = g) :=
eq.rec_on p (λg, !homotopy_equiv_eq) g
@ -155,7 +154,7 @@ namespace pi
/- Truncatedness: any dependent product of n-types is an n-type -/
open is_trunc
definition is_trunc_pi [instance] [H : funext.{l k}] (B : A → Type.{k}) (n : trunc_index)
definition is_trunc_pi [instance] (B : A → Type) (n : trunc_index)
[H : ∀a, is_trunc n (B a)] : is_trunc n (Πa, B a) :=
begin
reverts (B, H),
@ -175,7 +174,7 @@ namespace pi
is_trunc_eq n (f a) (g a)}
end
definition is_trunc_eq_pi [instance] [H : funext.{l k}] (n : trunc_index) (f g : Πa, B a)
definition is_trunc_eq_pi [instance] (n : trunc_index) (f g : Πa, B a)
[H : ∀a, is_trunc n (f a = g a)] : is_trunc n (f = g) :=
begin
apply is_trunc_equiv_closed, apply equiv.symm,

View file

@ -336,7 +336,6 @@ namespace sigma
/- *** The positive universal property. -/
section
open funext
definition is_equiv_sigma_rec [instance] (C : (Σa, B a) → Type)
: is_equiv (@sigma.rec _ _ C) :=
adjointify _ (λ g a b, g ⟨a, b⟩)

View file

@ -200,7 +200,7 @@ order for the change to take effect."
("eqn" . ,(lean-input-to-string-list "≠≁ ≉ ≄ ≇≆ ≢ ≭ "))
("=n" . (""))
("~" . ("")) ("~n" . (""))
("~" . ("")) ("~n" . ("")) ("homotopy" . (""))
("~~" . ("")) ("~~n" . (""))
("~~~" . (""))
(":~" . (""))
@ -317,10 +317,15 @@ order for the change to take effect."
("clr" . ("")) ("clR" . (""))
;; Various operators/symbols.
("tr" . (""))
("tr" . ,(lean-input-to-string-list "⬝▹"))
("con" . (""))
("cdot" . (""))
("sy" . ("⁻¹"))
("inv" . ("⁻¹"))
("-1" . ("⁻¹"))
("-1" . ("⁻¹"))
("-1p" . ("⁻¹ᵖ"))
("-1e" . ("⁻¹ᵉ"))
("-1h" . ("⁻¹ʰ"))
("qed" . (""))
("x" . ("×"))
("o" . (""))
@ -373,7 +378,7 @@ order for the change to take effect."
;; Arrows.
("l" . ,(lean-input-to-string-list "←⇐⇚⇇⇆↤⇦↞↼↽⇠⇺↜⇽⟵⟸↚⇍⇷ ↹ ↢↩↫⇋⇜⇤⟻⟽⤆↶↺⟲ "))
("l" . ,(lean-input-to-string-list "λ←⇐⇚⇇⇆↤⇦↞↼↽⇠⇺↜⇽⟵⟸↚⇍⇷ ↹ ↢↩↫⇋⇜⇤⟻⟽⤆↶↺⟲ "))
("r" . ,(lean-input-to-string-list "→⇒⇛⇉⇄↦⇨↠⇀⇁⇢⇻↝⇾⟶⟹↛⇏⇸⇶ ↴ ↣↪↬⇌⇝⇥⟼⟾⤇↷↻⟳⇰⇴⟴⟿ ➵➸➙➔➛➜➝➞➟➠➡➢➣➤➧➨➩➪➫➬➭➮➯➱➲➳➺➻➼➽➾⊸"))
("u" . ,(lean-input-to-string-list "↑⇑⟰⇈⇅↥⇧↟↿↾⇡⇞ ↰↱➦ ⇪⇫⇬⇭⇮⇯ "))
("d" . ,(lean-input-to-string-list "↓⇓⟱⇊⇵↧⇩↡⇃⇂⇣⇟ ↵↲↳➥ ↯ "))
@ -442,6 +447,7 @@ order for the change to take effect."
("t" . ,(lean-input-to-string-list "▸▹►▻◂◃◄◅▴▵▾▿◢◿◣◺◤◸◥◹"))
("Tr" . ,(lean-input-to-string-list "◀◁▶▷▲△▼▽◬◭◮"))
("transport" . (""))
("tb" . ,(lean-input-to-string-list "◂▸▴▾◄►◢◣◤◥"))
("tw" . ,(lean-input-to-string-list "◃▹▵▿◅▻◿◺◸◹"))