feat(library/hott): try to replace the proof irrelevance
This commit is contained in:
parent
5fe8ee606f
commit
d8e2206bbc
4 changed files with 53 additions and 5 deletions
|
@ -8,7 +8,8 @@
|
||||||
import .natural_transformation
|
import .natural_transformation
|
||||||
import data.unit data.sigma data.prod data.empty data.bool
|
import data.unit data.sigma data.prod data.empty data.bool
|
||||||
|
|
||||||
open eq eq.ops prod
|
open path prod
|
||||||
|
|
||||||
namespace category
|
namespace category
|
||||||
namespace opposite
|
namespace opposite
|
||||||
section
|
section
|
||||||
|
|
|
@ -3,7 +3,7 @@
|
||||||
-- Author: Floris van Doorn, Jakob von Raumer
|
-- Author: Floris van Doorn, Jakob von Raumer
|
||||||
|
|
||||||
import .functor hott.axioms.funext
|
import .functor hott.axioms.funext
|
||||||
open precategory path functor
|
open precategory path functor truncation
|
||||||
|
|
||||||
inductive natural_transformation {C D : Precategory} (F G : C ⇒ D) : Type :=
|
inductive natural_transformation {C D : Precategory} (F G : C ⇒ D) : Type :=
|
||||||
mk : Π (η : Π (a : C), hom (F a) (G a))
|
mk : Π (η : Π (a : C), hom (F a) (G a))
|
||||||
|
@ -35,20 +35,35 @@ namespace natural_transformation
|
||||||
|
|
||||||
infixr `∘n`:60 := compose
|
infixr `∘n`:60 := compose
|
||||||
|
|
||||||
definition foo (C : Precategory) (a b : C) (f g : hom a b) (p q : f ≈ g) : p ≈ q :=
|
/-definition foo (C : Precategory) (a b : C) (f g : hom a b) (p q : f ≈ g) : p ≈ q :=
|
||||||
@truncation.is_hset.elim _ !homH f g p q
|
@truncation.is_hset.elim _ !homH f g p q
|
||||||
|
|
||||||
|
definition nt_is_hset {F G : functor C D} : is_hset (F ⟹ G) := sorry
|
||||||
|
|
||||||
|
definition eta_nat_path {η₁ η₂ : F ⟹ G} (H1 : natural_map η₁ ≈ natural_map η₂)
|
||||||
|
(H2 : (H1 ▹ naturality η₁) ≈ naturality η₂) : η₁ ≈ η₂ :=
|
||||||
|
rec_on η₁ (λ eta1 nat1, rec_on η₂ (λ eta2 nat2, sorry))-/
|
||||||
|
|
||||||
protected definition assoc (η₃ : H ⟹ I) (η₂ : G ⟹ H) (η₁ : F ⟹ G) [fext : funext.{l_1 l_4}] :
|
protected definition assoc (η₃ : H ⟹ I) (η₂ : G ⟹ H) (η₁ : F ⟹ G) [fext : funext.{l_1 l_4}] :
|
||||||
η₃ ∘n (η₂ ∘n η₁) ≈ (η₃ ∘n η₂) ∘n η₁ :=
|
η₃ ∘n (η₂ ∘n η₁) ≈ (η₃ ∘n η₂) ∘n η₁ :=
|
||||||
dcongr_arg2 mk (funext.path_forall _ _ (λ x, !assoc)) sorry
|
have prir [visible] : is_hset (Π {a b : C} (f : hom a b), I f ∘ (η₃ ∘n η₂) a ∘ η₁ a ≈ ((η₃ ∘n η₂) b ∘ η₁ b) ∘ F f),
|
||||||
|
from sorry,
|
||||||
|
path.dcongr_arg2 mk
|
||||||
|
(funext.path_forall
|
||||||
|
(λ (x : C), η₃ x ∘ (η₂ x ∘ η₁ x))
|
||||||
|
(λ (x : C), (η₃ x ∘ η₂ x) ∘ η₁ x)
|
||||||
|
(λ x, assoc (η₃ x) (η₂ x) (η₁ x))
|
||||||
|
)
|
||||||
|
sorry --(@is_hset.elim _ _ _ _ _ _)
|
||||||
|
|
||||||
protected definition id {C D : Precategory} {F : functor C D} : natural_transformation F F :=
|
protected definition id {C D : Precategory} {F : functor C D} : natural_transformation 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) : natural_transformation F F := id
|
protected definition ID {C D : Precategory} (F : functor C D) : natural_transformation F F := id
|
||||||
|
|
||||||
protected theorem id_left (η : F ⟹ G) [fext : funext] : natural_transformation.compose id η ≈ η :=
|
protected theorem id_left (η : F ⟹ G) [fext : funext] : natural_transformation.compose id η ≈ η :=
|
||||||
rec (λf H, path.dcongr_arg2 mk (funext.path_forall _ _ (λ x, !id_left)) sorry) η
|
rec_on η (λf H, path.dcongr_arg2 mk (funext.path_forall _ _ (λ x, !id_left)) sorry)
|
||||||
|
|
||||||
protected theorem id_right (η : F ⟹ G) [fext : funext] : natural_transformation.compose η id ≈ η :=
|
protected theorem id_right (η : F ⟹ G) [fext : funext] : natural_transformation.compose η id ≈ η :=
|
||||||
rec (λf H, path.dcongr_arg2 mk (funext.path_forall _ _ (λ x, !id_right)) sorry) η
|
rec (λf H, path.dcongr_arg2 mk (funext.path_forall _ _ (λ x, !id_right)) sorry) η
|
||||||
|
|
||||||
end natural_transformation
|
end natural_transformation
|
||||||
|
|
|
@ -16,6 +16,7 @@ structure is_equiv [class] {A B : Type} (f : A → B) :=
|
||||||
(sect : (inv ∘ f) ∼ id)
|
(sect : (inv ∘ f) ∼ id)
|
||||||
(adj : Πx, retr (f x) ≈ ap f (sect x))
|
(adj : Πx, retr (f x) ≈ ap f (sect x))
|
||||||
|
|
||||||
|
|
||||||
-- A more bundled version of equivalence to calculate with
|
-- A more bundled version of equivalence to calculate with
|
||||||
structure equiv (A B : Type) :=
|
structure equiv (A B : Type) :=
|
||||||
(to_fun : A → B)
|
(to_fun : A → B)
|
||||||
|
|
|
@ -642,6 +642,36 @@ definition apD02_pp {A} (B : A → Type) (f : Π x:A, B x) {x y : A}
|
||||||
⬝ (whiskerR ((transport2_p2p B r1 r2 (f x))⁻¹) (apD f p3)) :=
|
⬝ (whiskerR ((transport2_p2p B r1 r2 (f x))⁻¹) (apD f p3)) :=
|
||||||
rec_on r2 (rec_on r1 (rec_on p1 idp))
|
rec_on r2 (rec_on r1 (rec_on p1 idp))
|
||||||
|
|
||||||
|
section
|
||||||
|
variables {A B C D E : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D}
|
||||||
|
|
||||||
|
theorem congr_arg2 (f : A → B → C) (Ha : a ≈ a') (Hb : b ≈ b') : f a b ≈ f a' b' :=
|
||||||
|
rec_on Ha (rec_on Hb idp)
|
||||||
|
|
||||||
|
theorem congr_arg3 (f : A → B → C → D) (Ha : a ≈ a') (Hb : b ≈ b') (Hc : c ≈ c')
|
||||||
|
: f a b c ≈ f a' b' c' :=
|
||||||
|
rec_on Ha (congr_arg2 (f a) Hb Hc)
|
||||||
|
|
||||||
|
theorem congr_arg4 (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' :=
|
||||||
|
rec_on Ha (congr_arg3 (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'}
|
||||||
|
|
||||||
|
theorem dcongr_arg2 (f : Πa, B a → F) (Ha : a ≈ a') (Hb : Ha ▹ b ≈ b')
|
||||||
|
: f a b ≈ f a' b' :=
|
||||||
|
rec_on Hb (rec_on Ha idp)
|
||||||
|
|
||||||
|
end
|
||||||
|
|
||||||
/- From the Coq version:
|
/- From the Coq version:
|
||||||
|
|
||||||
-- ** Tactics, hints, and aliases
|
-- ** Tactics, hints, and aliases
|
||||||
|
@ -694,4 +724,5 @@ Ltac hott_simpl :=
|
||||||
autorewrite with paths in * |- * ; auto with path_hints.
|
autorewrite with paths in * |- * ; auto with path_hints.
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
end path
|
end path
|
||||||
|
|
Loading…
Reference in a new issue