refactor(hott,library): test new tactics in the HoTT and standard libraries
This commit is contained in:
parent
657ad3327f
commit
4e1146a2d5
8 changed files with 104 additions and 102 deletions
|
@ -39,10 +39,9 @@ namespace category
|
||||||
(λc d f,
|
(λc d f,
|
||||||
begin
|
begin
|
||||||
apply comp_inverse_eq_of_eq_comp,
|
apply comp_inverse_eq_of_eq_comp,
|
||||||
apply concat, rotate_left 1, apply assoc,
|
transitivity (natural_map η d)⁻¹ ∘ to_fun_hom G f ∘ natural_map η c,
|
||||||
apply eq_inverse_comp_of_comp_eq,
|
{apply eq_inverse_comp_of_comp_eq, symmetry, apply naturality},
|
||||||
apply inverse,
|
{apply assoc}
|
||||||
apply naturality,
|
|
||||||
end)
|
end)
|
||||||
|
|
||||||
definition nat_trans_left_inverse : nat_trans_inverse η ∘n η = nat_trans.id :=
|
definition nat_trans_left_inverse : nat_trans_inverse η ∘n η = nat_trans.id :=
|
||||||
|
@ -80,14 +79,14 @@ namespace category
|
||||||
|
|
||||||
definition naturality_iso {c c' : C} (f : c ⟶ c') : G f = η c' ∘ F f ∘ (η c)⁻¹ :=
|
definition naturality_iso {c c' : C} (f : c ⟶ c') : G f = η c' ∘ F f ∘ (η c)⁻¹ :=
|
||||||
calc
|
calc
|
||||||
G f = (G f ∘ η c) ∘ (η c)⁻¹ : comp_inverse_cancel_right
|
G f = (G f ∘ η c) ∘ (η c)⁻¹ : by rewrite comp_inverse_cancel_right
|
||||||
... = (η c' ∘ F f) ∘ (η c)⁻¹ : by rewrite naturality
|
... = (η c' ∘ F f) ∘ (η c)⁻¹ : by rewrite naturality
|
||||||
... = η c' ∘ F f ∘ (η c)⁻¹ : assoc
|
... = η c' ∘ F f ∘ (η c)⁻¹ : by rewrite assoc
|
||||||
|
|
||||||
definition naturality_iso' {c c' : C} (f : c ⟶ c') : (η c')⁻¹ ∘ G f ∘ η c = F f :=
|
definition naturality_iso' {c c' : C} (f : c ⟶ c') : (η c')⁻¹ ∘ G f ∘ η c = F f :=
|
||||||
calc
|
calc
|
||||||
(η c')⁻¹ ∘ G f ∘ η c = (η c')⁻¹ ∘ η c' ∘ F f : by rewrite naturality
|
(η c')⁻¹ ∘ G f ∘ η c = (η c')⁻¹ ∘ η c' ∘ F f : by rewrite naturality
|
||||||
... = F f : inverse_comp_cancel_left
|
... = F f : by rewrite inverse_comp_cancel_left
|
||||||
|
|
||||||
omit isoη
|
omit isoη
|
||||||
|
|
||||||
|
@ -128,7 +127,7 @@ namespace category
|
||||||
{apply (ap (λx, to_hom x ∘ to_fun_hom F f ∘ _)), apply (right_inv iso_of_eq)},
|
{apply (ap (λx, to_hom x ∘ to_fun_hom F f ∘ _)), apply (right_inv iso_of_eq)},
|
||||||
apply concat,
|
apply concat,
|
||||||
{apply (ap (λx, _ ∘ to_fun_hom F f ∘ (to_hom x)⁻¹)), apply (right_inv iso_of_eq)},
|
{apply (ap (λx, _ ∘ to_fun_hom F f ∘ (to_hom x)⁻¹)), apply (right_inv iso_of_eq)},
|
||||||
apply inverse, apply naturality_iso}
|
symmetry, apply naturality_iso}
|
||||||
end
|
end
|
||||||
|
|
||||||
definition iso_of_eq_eq_of_iso (η : F ≅ G) : iso_of_eq (eq_of_iso η) = η :=
|
definition iso_of_eq_eq_of_iso (η : F ≅ G) : iso_of_eq (eq_of_iso η) = η :=
|
||||||
|
|
|
@ -74,7 +74,7 @@ namespace functor
|
||||||
apply (apd10' c'),
|
apply (apd10' c'),
|
||||||
apply concat, rotate_left 1, esimp,
|
apply concat, rotate_left 1, esimp,
|
||||||
exact (pi_transport_constant (eq_of_homotopy pF) H₁ c),
|
exact (pi_transport_constant (eq_of_homotopy pF) H₁ c),
|
||||||
apply idp
|
reflexivity
|
||||||
end))))
|
end))))
|
||||||
|
|
||||||
definition functor_eq {F₁ F₂ : C ⇒ D} : Π(p : to_fun_ob F₁ ∼ to_fun_ob F₂),
|
definition functor_eq {F₁ F₂ : C ⇒ D} : Π(p : to_fun_ob F₁ ∼ to_fun_ob F₂),
|
||||||
|
@ -91,7 +91,7 @@ namespace functor
|
||||||
is_iso (F f) :=
|
is_iso (F f) :=
|
||||||
begin
|
begin
|
||||||
fapply @is_iso.mk, apply (F (f⁻¹)),
|
fapply @is_iso.mk, apply (F (f⁻¹)),
|
||||||
repeat (apply concat ; apply inverse ; apply (respect_comp F) ;
|
repeat (apply concat ; symmetry ; apply (respect_comp F) ;
|
||||||
apply concat ; apply (ap (λ x, to_fun_hom F x)) ;
|
apply concat ; apply (ap (λ x, to_fun_hom F x)) ;
|
||||||
(apply left_inverse | apply right_inverse);
|
(apply left_inverse | apply right_inverse);
|
||||||
apply (respect_id F) ),
|
apply (respect_id F) ),
|
||||||
|
@ -104,10 +104,12 @@ namespace functor
|
||||||
F (f⁻¹) = (F f)⁻¹ :=
|
F (f⁻¹) = (F f)⁻¹ :=
|
||||||
begin
|
begin
|
||||||
fapply @left_inverse_eq_right_inverse, apply (F f),
|
fapply @left_inverse_eq_right_inverse, apply (F f),
|
||||||
apply concat, apply inverse, apply (respect_comp F),
|
transitivity to_fun_hom F (f⁻¹ ∘ f),
|
||||||
apply concat, apply (ap (λ x, to_fun_hom F x)),
|
{symmetry, apply (respect_comp F)},
|
||||||
apply left_inverse, apply respect_id,
|
{transitivity to_fun_hom F category.id,
|
||||||
apply right_inverse,
|
{congruence, apply left_inverse},
|
||||||
|
{apply respect_id}},
|
||||||
|
apply right_inverse
|
||||||
end
|
end
|
||||||
|
|
||||||
protected definition assoc (H : C ⇒ D) (G : B ⇒ C) (F : A ⇒ B) :
|
protected definition assoc (H : C ⇒ D) (G : B ⇒ C) (F : A ⇒ B) :
|
||||||
|
@ -141,12 +143,12 @@ namespace functor
|
||||||
exact ⟨d1, d2, (d3, @d4)⟩},
|
exact ⟨d1, d2, (d3, @d4)⟩},
|
||||||
{intro F,
|
{intro F,
|
||||||
cases F,
|
cases F,
|
||||||
apply idp},
|
reflexivity},
|
||||||
{intro S,
|
{intro S,
|
||||||
cases S with d1 S2,
|
cases S with d1 S2,
|
||||||
cases S2 with d2 P1,
|
cases S2 with d2 P1,
|
||||||
cases P1,
|
cases P1,
|
||||||
apply idp},
|
reflexivity},
|
||||||
end
|
end
|
||||||
|
|
||||||
section
|
section
|
||||||
|
@ -172,8 +174,7 @@ namespace functor
|
||||||
begin
|
begin
|
||||||
cases p, cases F₁,
|
cases p, cases F₁,
|
||||||
apply concat, rotate_left 1, apply functor_eq'_idp,
|
apply concat, rotate_left 1, apply functor_eq'_idp,
|
||||||
apply (ap (functor_eq' idp)),
|
esimp
|
||||||
apply idp_con,
|
|
||||||
end
|
end
|
||||||
|
|
||||||
definition functor_eq2' {F₁ F₂ : C ⇒ D} {p₁ p₂ : to_fun_ob F₁ = to_fun_ob F₂} (q₁ q₂)
|
definition functor_eq2' {F₁ F₂ : C ⇒ D} {p₁ p₂ : to_fun_ob F₁ = to_fun_ob F₂} (q₁ q₂)
|
||||||
|
|
|
@ -129,29 +129,30 @@ namespace functor
|
||||||
begin
|
begin
|
||||||
intro cd cd' fg,
|
intro cd cd' fg,
|
||||||
cases cd with c d, cases cd' with c' d', cases fg with f g,
|
cases cd with c d, cases cd' with c' d', cases fg with f g,
|
||||||
apply concat, apply id_leftright,
|
transitivity to_fun_hom (functor_uncurry (functor_curry F)) (f, g),
|
||||||
|
apply id_leftright,
|
||||||
show (functor_uncurry (functor_curry F)) (f, g) = F (f,g),
|
show (functor_uncurry (functor_curry F)) (f, g) = F (f,g),
|
||||||
from calc
|
from calc
|
||||||
(functor_uncurry (functor_curry F)) (f, g) = to_fun_hom F (id, g) ∘ to_fun_hom F (f, id) : by esimp
|
(functor_uncurry (functor_curry F)) (f, g) = to_fun_hom F (id, g) ∘ to_fun_hom F (f, id) : by esimp
|
||||||
... = F (id ∘ f, g ∘ id) : respect_comp F (id,g) (f,id)
|
... = F (id ∘ f, g ∘ id) : respect_comp F (id,g) (f,id)
|
||||||
... = F (f, g ∘ id) : by rewrite id_left
|
... = F (f, g ∘ id) : by rewrite id_left
|
||||||
... = F (f,g) : by rewrite id_right,
|
... = F (f,g) : by rewrite id_right,
|
||||||
end
|
end
|
||||||
|
|
||||||
definition functor_curry_functor_uncurry_ob (c : C)
|
definition functor_curry_functor_uncurry_ob (c : C)
|
||||||
: functor_curry (functor_uncurry G) c = G c :=
|
: functor_curry (functor_uncurry G) c = G c :=
|
||||||
begin
|
begin
|
||||||
fapply functor_eq,
|
fapply functor_eq,
|
||||||
{intro d, apply idp},
|
{intro d, reflexivity},
|
||||||
{intro d d' g,
|
{intro d d' g,
|
||||||
apply concat, apply id_leftright,
|
apply concat, apply id_leftright,
|
||||||
show to_fun_hom (functor_curry (functor_uncurry G) c) g = to_fun_hom (G c) g,
|
show to_fun_hom (functor_curry (functor_uncurry G) c) g = to_fun_hom (G c) g,
|
||||||
from calc
|
from calc
|
||||||
to_fun_hom (functor_curry (functor_uncurry G) c) g
|
to_fun_hom (functor_curry (functor_uncurry G) c) g
|
||||||
= to_fun_hom (G c) g ∘ natural_map (to_fun_hom G (ID c)) d : by esimp
|
= to_fun_hom (G c) g ∘ natural_map (to_fun_hom G (ID c)) d : by esimp
|
||||||
... = to_fun_hom (G c) g ∘ natural_map (ID (G c)) d : by rewrite respect_id
|
... = to_fun_hom (G c) g ∘ natural_map (ID (G c)) d : by rewrite respect_id
|
||||||
... = to_fun_hom (G c) g ∘ id : idp
|
... = to_fun_hom (G c) g ∘ id : by reflexivity
|
||||||
... = to_fun_hom (G c) g : id_right}
|
... = to_fun_hom (G c) g : id_right}
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem functor_curry_functor_uncurry : functor_curry (functor_uncurry G) = G :=
|
theorem functor_curry_functor_uncurry : functor_curry (functor_uncurry G) = G :=
|
||||||
|
@ -163,7 +164,7 @@ namespace functor
|
||||||
apply concat,
|
apply concat,
|
||||||
{apply (ap (λx, x ∘ _)),
|
{apply (ap (λx, x ∘ _)),
|
||||||
apply concat, apply natural_map_hom_of_eq, apply (ap hom_of_eq), apply ap010_functor_eq},
|
apply concat, apply natural_map_hom_of_eq, apply (ap hom_of_eq), apply ap010_functor_eq},
|
||||||
apply concat,
|
apply concat,
|
||||||
{apply (ap (λx, _ ∘ x)), apply (ap (λx, _ ∘ x)),
|
{apply (ap (λx, _ ∘ x)), apply (ap (λx, _ ∘ x)),
|
||||||
apply concat, apply natural_map_inv_of_eq,
|
apply concat, apply natural_map_inv_of_eq,
|
||||||
apply (ap (λx, hom_of_eq x⁻¹)), apply ap010_functor_eq},
|
apply (ap (λx, hom_of_eq x⁻¹)), apply ap010_functor_eq},
|
||||||
|
|
|
@ -50,54 +50,54 @@ namespace eq
|
||||||
notation f `∼3`:50 g := homotopy3 f g
|
notation f `∼3`:50 g := homotopy3 f g
|
||||||
|
|
||||||
definition ap011 (f : U → V → W) (Hu : u = u') (Hv : v = v') : f u v = f u' v' :=
|
definition ap011 (f : U → V → W) (Hu : u = u') (Hv : v = v') : f u v = f u' v' :=
|
||||||
by cases Hu; exact (ap (f u) Hv)
|
by cases Hu; congruence; repeat assumption
|
||||||
|
|
||||||
definition ap0111 (f : U → V → W → X) (Hu : u = u') (Hv : v = v') (Hw : w = w')
|
definition ap0111 (f : U → V → W → X) (Hu : u = u') (Hv : v = v') (Hw : w = w')
|
||||||
: f u v w = f u' v' w' :=
|
: f u v w = f u' v' w' :=
|
||||||
by cases Hu; exact (ap011 (f u) Hv Hw)
|
by cases Hu; congruence; repeat assumption
|
||||||
|
|
||||||
definition ap01111 (f : U → V → W → X → Y) (Hu : u = u') (Hv : v = v') (Hw : w = w') (Hx : x = x')
|
definition ap01111 (f : U → V → W → X → Y) (Hu : u = u') (Hv : v = v') (Hw : w = w') (Hx : x = x')
|
||||||
: f u v w x = f u' v' w' x' :=
|
: f u v w x = f u' v' w' x' :=
|
||||||
by cases Hu; exact (ap0111 (f u) Hv Hw Hx)
|
by cases Hu; congruence; repeat assumption
|
||||||
|
|
||||||
definition ap011111 (f : U → V → W → X → Y → Z)
|
definition ap011111 (f : U → V → W → X → Y → Z)
|
||||||
(Hu : u = u') (Hv : v = v') (Hw : w = w') (Hx : x = x') (Hy : y = y')
|
(Hu : u = u') (Hv : v = v') (Hw : w = w') (Hx : x = x') (Hy : y = y')
|
||||||
: f u v w x y = f u' v' w' x' y' :=
|
: f u v w x y = f u' v' w' x' y' :=
|
||||||
by cases Hu; exact (ap01111 (f u) Hv Hw Hx Hy)
|
by cases Hu; congruence; repeat assumption
|
||||||
|
|
||||||
definition ap0111111 (f : U → V → W → X → Y → Z → A)
|
definition ap0111111 (f : U → V → W → X → Y → Z → A)
|
||||||
(Hu : u = u') (Hv : v = v') (Hw : w = w') (Hx : x = x') (Hy : y = y') (Hz : z = z')
|
(Hu : u = u') (Hv : v = v') (Hw : w = w') (Hx : x = x') (Hy : y = y') (Hz : z = z')
|
||||||
: f u v w x y z = f u' v' w' x' y' z' :=
|
: f u v w x y z = f u' v' w' x' y' z' :=
|
||||||
by cases Hu; exact (ap011111 (f u) Hv Hw Hx Hy Hz)
|
by cases Hu; congruence; repeat assumption
|
||||||
|
|
||||||
definition ap010 (f : X → Πa, B a) (Hx : x = x') : f x ∼ f x' :=
|
definition ap010 (f : X → Πa, B a) (Hx : x = x') : f x ∼ f x' :=
|
||||||
λa, by cases Hx; exact idp
|
by intros; cases Hx; reflexivity
|
||||||
|
|
||||||
definition ap0100 (f : X → Πa b, C a b) (Hx : x = x') : f x ∼2 f x' :=
|
definition ap0100 (f : X → Πa b, C a b) (Hx : x = x') : f x ∼2 f x' :=
|
||||||
λa b, by cases Hx; exact idp
|
by intros; cases Hx; reflexivity
|
||||||
|
|
||||||
definition ap01000 (f : X → Πa b c, D a b c) (Hx : x = x') : f x ∼3 f x' :=
|
definition ap01000 (f : X → Πa b c, D a b c) (Hx : x = x') : f x ∼3 f x' :=
|
||||||
λa b c, by cases Hx; exact idp
|
by intros; cases Hx; reflexivity
|
||||||
|
|
||||||
definition apd011 (f : Πa, B a → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
definition apd011 (f : Πa, B a → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
||||||
: f a b = f a' b' :=
|
: f a b = f a' b' :=
|
||||||
by cases Ha; cases Hb; exact idp
|
by cases Ha; cases Hb; reflexivity
|
||||||
|
|
||||||
definition apd0111 (f : Πa b, C a b → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
definition apd0111 (f : Πa b, C a b → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
||||||
(Hc : apd011 C Ha Hb ▹ c = c')
|
(Hc : apd011 C Ha Hb ▹ c = c')
|
||||||
: f a b c = f a' b' c' :=
|
: f a b c = f a' b' c' :=
|
||||||
by cases Ha; cases Hb; cases Hc; exact idp
|
by cases Ha; cases Hb; cases Hc; reflexivity
|
||||||
|
|
||||||
definition apd01111 (f : Πa b c, D a b c → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
definition apd01111 (f : Πa b c, D a b c → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
||||||
(Hc : apd011 C Ha Hb ▹ c = c') (Hd : apd0111 D Ha Hb Hc ▹ d = d')
|
(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' :=
|
: f a b c d = f a' b' c' d' :=
|
||||||
by cases Ha; cases Hb; cases Hc; cases Hd; exact idp
|
by cases Ha; cases Hb; cases Hc; cases Hd; reflexivity
|
||||||
|
|
||||||
definition apd011111 (f : Πa b c d, E a b c d → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
definition apd011111 (f : Πa b c d, E a b c d → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
||||||
(Hc : apd011 C Ha Hb ▹ c = c') (Hd : apd0111 D Ha Hb Hc ▹ d = d')
|
(Hc : apd011 C Ha Hb ▹ c = c') (Hd : apd0111 D Ha Hb Hc ▹ d = d')
|
||||||
(He : apd01111 E Ha Hb Hc Hd ▹ e = e')
|
(He : apd01111 E Ha Hb Hc Hd ▹ e = e')
|
||||||
: f a b c d e = f a' b' c' d' e' :=
|
: f a b c d e = f a' b' c' d' e' :=
|
||||||
by cases Ha; cases Hb; cases Hc; cases Hd; cases He; exact idp
|
by cases Ha; cases Hb; cases Hc; cases Hd; cases He; reflexivity
|
||||||
|
|
||||||
definition apd100 {f g : Πa b, C a b} (p : f = g) : f ∼2 g :=
|
definition apd100 {f g : Πa b, C a b} (p : f = g) : f ∼2 g :=
|
||||||
λa b, apd10 (apd10 p a) b
|
λa b, apd10 (apd10 p a) b
|
||||||
|
@ -128,17 +128,17 @@ namespace eq
|
||||||
definition eq_of_homotopy2_id (f : Πa b, C a b)
|
definition eq_of_homotopy2_id (f : Πa b, C a b)
|
||||||
: eq_of_homotopy2 (λa b, idpath (f a b)) = idpath f :=
|
: eq_of_homotopy2 (λa b, idpath (f a b)) = idpath f :=
|
||||||
begin
|
begin
|
||||||
apply concat,
|
transitivity eq_of_homotopy (λ a, idpath (f a)),
|
||||||
{apply (ap (λx, eq_of_homotopy x)), apply eq_of_homotopy, intro a, apply eq_of_homotopy_idp},
|
{apply (ap eq_of_homotopy), apply eq_of_homotopy, intros, apply eq_of_homotopy_idp},
|
||||||
apply eq_of_homotopy_idp
|
apply eq_of_homotopy_idp
|
||||||
end
|
end
|
||||||
|
|
||||||
definition eq_of_homotopy3_id (f : Πa b c, D a b c)
|
definition eq_of_homotopy3_id (f : Πa b c, D a b c)
|
||||||
: eq_of_homotopy3 (λa b c, idpath (f a b c)) = idpath f :=
|
: eq_of_homotopy3 (λa b c, idpath (f a b c)) = idpath f :=
|
||||||
begin
|
begin
|
||||||
apply concat,
|
transitivity _,
|
||||||
{apply (ap (λx, eq_of_homotopy x)), apply eq_of_homotopy, intro a, apply eq_of_homotopy2_id},
|
{apply (ap eq_of_homotopy), apply eq_of_homotopy, intros, apply eq_of_homotopy2_id},
|
||||||
apply eq_of_homotopy_idp
|
apply eq_of_homotopy_idp
|
||||||
end
|
end
|
||||||
|
|
||||||
end eq
|
end eq
|
||||||
|
|
|
@ -35,8 +35,8 @@ parameters {A : Type} (R : A → A → hprop)
|
||||||
{ intro x', apply Pt},
|
{ intro x', apply Pt},
|
||||||
{ intro y, fapply (type_quotient.rec_on y),
|
{ intro y, fapply (type_quotient.rec_on y),
|
||||||
{ exact Pc},
|
{ exact Pc},
|
||||||
{ intro a a' H,
|
{ intros,
|
||||||
apply concat, apply transport_compose;apply Pp}}
|
apply concat, apply transport_compose; apply Pp}}
|
||||||
end
|
end
|
||||||
|
|
||||||
protected definition rec_on [reducible] {P : quotient → Type} (x : quotient)
|
protected definition rec_on [reducible] {P : quotient → Type} (x : quotient)
|
||||||
|
|
|
@ -62,19 +62,19 @@ namespace eq
|
||||||
|
|
||||||
definition transport_eq_l (p : a1 = a2) (q : a1 = a3)
|
definition transport_eq_l (p : a1 = a2) (q : a1 = a3)
|
||||||
: transport (λx, x = a3) p q = p⁻¹ ⬝ q :=
|
: transport (λx, x = a3) p q = p⁻¹ ⬝ q :=
|
||||||
by cases p; cases q; apply idp
|
by cases p; cases q; reflexivity
|
||||||
|
|
||||||
definition transport_eq_r (p : a2 = a3) (q : a1 = a2)
|
definition transport_eq_r (p : a2 = a3) (q : a1 = a2)
|
||||||
: transport (λx, a1 = x) p q = q ⬝ p :=
|
: transport (λx, a1 = x) p q = q ⬝ p :=
|
||||||
by cases p; cases q; apply idp
|
by cases p; cases q; reflexivity
|
||||||
|
|
||||||
definition transport_eq_lr (p : a1 = a2) (q : a1 = a1)
|
definition transport_eq_lr (p : a1 = a2) (q : a1 = a1)
|
||||||
: transport (λx, x = x) p q = p⁻¹ ⬝ q ⬝ p :=
|
: transport (λx, x = x) p q = p⁻¹ ⬝ q ⬝ p :=
|
||||||
begin
|
begin
|
||||||
apply (eq.rec_on p),
|
cases p,
|
||||||
apply inverse, apply concat,
|
symmetry, transitivity (refl a1)⁻¹ ⬝ q,
|
||||||
apply con_idp,
|
apply con_idp,
|
||||||
apply idp_con
|
apply idp_con
|
||||||
end
|
end
|
||||||
|
|
||||||
definition transport_eq_Fl (p : a1 = a2) (q : f a1 = b)
|
definition transport_eq_Fl (p : a1 = a2) (q : f a1 = b)
|
||||||
|
@ -88,42 +88,44 @@ namespace eq
|
||||||
definition transport_eq_FlFr (p : a1 = a2) (q : f a1 = g a1)
|
definition transport_eq_FlFr (p : a1 = a2) (q : f a1 = g a1)
|
||||||
: transport (λx, f x = g x) p q = (ap f p)⁻¹ ⬝ q ⬝ (ap g p) :=
|
: transport (λx, f x = g x) p q = (ap f p)⁻¹ ⬝ q ⬝ (ap g p) :=
|
||||||
begin
|
begin
|
||||||
apply (eq.rec_on p),
|
cases p,
|
||||||
apply inverse, apply concat,
|
symmetry, transitivity (ap f (refl a1))⁻¹ ⬝ q,
|
||||||
apply con_idp,
|
apply con_idp,
|
||||||
apply idp_con
|
apply idp_con
|
||||||
end
|
end
|
||||||
|
|
||||||
definition transport_eq_FlFr_D {B : A → Type} {f g : Πa, B a}
|
definition transport_eq_FlFr_D {B : A → Type} {f g : Πa, B a}
|
||||||
(p : a1 = a2) (q : f a1 = g a1)
|
(p : a1 = a2) (q : f a1 = g a1)
|
||||||
: transport (λx, f x = g x) p q = (apd f p)⁻¹ ⬝ ap (transport B p) q ⬝ (apd g p) :=
|
: transport (λx, f x = g x) p q = (apd f p)⁻¹ ⬝ ap (transport B p) q ⬝ (apd g p) :=
|
||||||
begin
|
begin
|
||||||
apply (eq.rec_on p),
|
cases p,
|
||||||
apply inverse,
|
symmetry,
|
||||||
apply concat, apply con_idp,
|
transitivity _,
|
||||||
apply concat, apply idp_con,
|
apply con_idp,
|
||||||
apply ap_id
|
transitivity _,
|
||||||
|
apply idp_con,
|
||||||
|
apply ap_id
|
||||||
end
|
end
|
||||||
|
|
||||||
definition transport_eq_FFlr (p : a1 = a2) (q : h (f a1) = a1)
|
definition transport_eq_FFlr (p : a1 = a2) (q : h (f a1) = a1)
|
||||||
: transport (λx, h (f x) = x) p q = (ap h (ap f p))⁻¹ ⬝ q ⬝ p :=
|
: transport (λx, h (f x) = x) p q = (ap h (ap f p))⁻¹ ⬝ q ⬝ p :=
|
||||||
begin
|
begin
|
||||||
apply (eq.rec_on p),
|
cases p,
|
||||||
apply inverse,
|
symmetry,
|
||||||
apply concat, apply con_idp,
|
transitivity (ap h (ap f (refl a1)))⁻¹ ⬝ q,
|
||||||
apply idp_con,
|
apply con_idp,
|
||||||
|
apply idp_con,
|
||||||
end
|
end
|
||||||
|
|
||||||
definition transport_eq_lFFr (p : a1 = a2) (q : a1 = h (f a1))
|
definition transport_eq_lFFr (p : a1 = a2) (q : a1 = h (f a1))
|
||||||
: transport (λx, x = h (f x)) p q = p⁻¹ ⬝ q ⬝ (ap h (ap f p)) :=
|
: transport (λx, x = h (f x)) p q = p⁻¹ ⬝ q ⬝ (ap h (ap f p)) :=
|
||||||
begin
|
begin
|
||||||
apply (eq.rec_on p),
|
cases p, symmetry,
|
||||||
apply inverse,
|
transitivity (refl a1)⁻¹ ⬝ q,
|
||||||
apply concat, apply con_idp,
|
apply con_idp,
|
||||||
apply idp_con,
|
apply idp_con,
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
||||||
-- The Functorial action of paths is [ap].
|
-- The Functorial action of paths is [ap].
|
||||||
|
|
||||||
/- Equivalences between path spaces -/
|
/- Equivalences between path spaces -/
|
||||||
|
@ -193,8 +195,8 @@ namespace eq
|
||||||
begin
|
begin
|
||||||
fapply adjointify,
|
fapply adjointify,
|
||||||
{intro s, apply (!cancel_right s)},
|
{intro s, apply (!cancel_right s)},
|
||||||
{intro s, cases r, cases s, cases q, apply idp},
|
{intro s, cases r, cases s, cases q, reflexivity},
|
||||||
{intro s, cases s, cases r, cases p, apply idp}
|
{intro s, cases s, cases r, cases p, reflexivity}
|
||||||
end
|
end
|
||||||
|
|
||||||
definition eq_equiv_con_eq_con_right (p q : a1 = a2) (r : a2 = a3) : (p = q) ≃ (p ⬝ r = q ⬝ r) :=
|
definition eq_equiv_con_eq_con_right (p q : a1 = a2) (r : a2 = a3) : (p = q) ≃ (p ⬝ r = q ⬝ r) :=
|
||||||
|
|
|
@ -29,7 +29,7 @@ namespace sigma
|
||||||
| eta3 ⟨u₁, u₂, u₃, u₄⟩ := idp
|
| eta3 ⟨u₁, u₂, u₃, u₄⟩ := idp
|
||||||
|
|
||||||
definition dpair_eq_dpair (p : a = a') (q : p ▹ b = b') : ⟨a, b⟩ = ⟨a', b'⟩ :=
|
definition dpair_eq_dpair (p : a = a') (q : p ▹ b = b') : ⟨a, b⟩ = ⟨a', b'⟩ :=
|
||||||
by cases p; cases q; apply idp
|
by cases p; cases q; reflexivity
|
||||||
|
|
||||||
definition sigma_eq (p : u.1 = v.1) (q : p ▹ u.2 = v.2) : u = v :=
|
definition sigma_eq (p : u.1 = v.1) (q : p ▹ u.2 = v.2) : u = v :=
|
||||||
by cases u; cases v; apply (dpair_eq_dpair p q)
|
by cases u; cases v; apply (dpair_eq_dpair p q)
|
||||||
|
@ -42,13 +42,13 @@ namespace sigma
|
||||||
postfix `..1`:(max+1) := eq_pr1
|
postfix `..1`:(max+1) := eq_pr1
|
||||||
|
|
||||||
definition eq_pr2 (p : u = v) : p..1 ▹ u.2 = v.2 :=
|
definition eq_pr2 (p : u = v) : p..1 ▹ u.2 = v.2 :=
|
||||||
by cases p; apply idp
|
by cases p; reflexivity
|
||||||
|
|
||||||
postfix `..2`:(max+1) := eq_pr2
|
postfix `..2`:(max+1) := eq_pr2
|
||||||
|
|
||||||
private definition dpair_sigma_eq (p : u.1 = v.1) (q : p ▹ u.2 = v.2)
|
private definition dpair_sigma_eq (p : u.1 = v.1) (q : p ▹ u.2 = v.2)
|
||||||
: ⟨(sigma_eq p q)..1, (sigma_eq p q)..2⟩ = ⟨p, q⟩ :=
|
: ⟨(sigma_eq p q)..1, (sigma_eq p q)..2⟩ = ⟨p, q⟩ :=
|
||||||
by cases u; cases v; cases p; cases q; apply idp
|
by cases u; cases v; cases p; cases q; reflexivity
|
||||||
|
|
||||||
definition sigma_eq_pr1 (p : u.1 = v.1) (q : p ▹ u.2 = v.2) : (sigma_eq p q)..1 = p :=
|
definition sigma_eq_pr1 (p : u.1 = v.1) (q : p ▹ u.2 = v.2) : (sigma_eq p q)..1 = p :=
|
||||||
(dpair_sigma_eq p q)..1
|
(dpair_sigma_eq p q)..1
|
||||||
|
@ -58,11 +58,11 @@ namespace sigma
|
||||||
(dpair_sigma_eq p q)..2
|
(dpair_sigma_eq p q)..2
|
||||||
|
|
||||||
definition sigma_eq_eta (p : u = v) : sigma_eq (p..1) (p..2) = p :=
|
definition sigma_eq_eta (p : u = v) : sigma_eq (p..1) (p..2) = p :=
|
||||||
by cases p; cases u; apply idp
|
by cases p; cases u; reflexivity
|
||||||
|
|
||||||
definition tr_pr1_sigma_eq {B' : A → Type} (p : u.1 = v.1) (q : p ▹ u.2 = v.2)
|
definition tr_pr1_sigma_eq {B' : A → Type} (p : u.1 = v.1) (q : p ▹ u.2 = v.2)
|
||||||
: transport (λx, B' x.1) (sigma_eq p q) = transport B' p :=
|
: transport (λx, B' x.1) (sigma_eq p q) = transport B' p :=
|
||||||
by cases u; cases v; cases p; cases q; apply idp
|
by cases u; cases v; cases p; cases q; reflexivity
|
||||||
|
|
||||||
/- the uncurried version of sigma_eq. We will prove that this is an equivalence -/
|
/- the uncurried version of sigma_eq. We will prove that this is an equivalence -/
|
||||||
|
|
||||||
|
@ -103,7 +103,7 @@ namespace sigma
|
||||||
(p2 : a' = a'') (q2 : p2 ▹ b' = b'') :
|
(p2 : a' = a'') (q2 : p2 ▹ b' = b'') :
|
||||||
dpair_eq_dpair (p1 ⬝ p2) (con_tr p1 p2 b ⬝ ap (transport B p2) q1 ⬝ q2)
|
dpair_eq_dpair (p1 ⬝ p2) (con_tr p1 p2 b ⬝ ap (transport B p2) q1 ⬝ q2)
|
||||||
= dpair_eq_dpair p1 q1 ⬝ dpair_eq_dpair p2 q2 :=
|
= dpair_eq_dpair p1 q1 ⬝ dpair_eq_dpair p2 q2 :=
|
||||||
by cases p1; cases p2; cases q1; cases q2; apply idp
|
by cases p1; cases p2; cases q1; cases q2; reflexivity
|
||||||
|
|
||||||
definition sigma_eq_con (p1 : u.1 = v.1) (q1 : p1 ▹ u.2 = v.2)
|
definition sigma_eq_con (p1 : u.1 = v.1) (q1 : p1 ▹ u.2 = v.2)
|
||||||
(p2 : v.1 = w.1) (q2 : p2 ▹ v.2 = w.2) :
|
(p2 : v.1 = w.1) (q2 : p2 ▹ v.2 = w.2) :
|
||||||
|
@ -114,7 +114,7 @@ namespace sigma
|
||||||
local attribute dpair_eq_dpair [reducible]
|
local attribute dpair_eq_dpair [reducible]
|
||||||
definition dpair_eq_dpair_con_idp (p : a = a') (q : p ▹ b = b') :
|
definition dpair_eq_dpair_con_idp (p : a = a') (q : p ▹ b = b') :
|
||||||
dpair_eq_dpair p q = dpair_eq_dpair p idp ⬝ dpair_eq_dpair idp q :=
|
dpair_eq_dpair p q = dpair_eq_dpair p idp ⬝ dpair_eq_dpair idp q :=
|
||||||
by cases p; cases q; apply idp
|
by cases p; cases q; reflexivity
|
||||||
|
|
||||||
/- eq_pr1 commutes with the groupoid structure. -/
|
/- eq_pr1 commutes with the groupoid structure. -/
|
||||||
|
|
||||||
|
@ -125,17 +125,17 @@ namespace sigma
|
||||||
/- Applying dpair to one argument is the same as dpair_eq_dpair with reflexivity in the first place. -/
|
/- Applying dpair to one argument is the same as dpair_eq_dpair with reflexivity in the first place. -/
|
||||||
|
|
||||||
definition ap_dpair (q : b₁ = b₂) : ap (sigma.mk a) q = dpair_eq_dpair idp q :=
|
definition ap_dpair (q : b₁ = b₂) : ap (sigma.mk a) q = dpair_eq_dpair idp q :=
|
||||||
by cases q; apply idp
|
by cases q; reflexivity
|
||||||
|
|
||||||
/- Dependent transport is the same as transport along a sigma_eq. -/
|
/- Dependent transport is the same as transport along a sigma_eq. -/
|
||||||
|
|
||||||
definition transportD_eq_transport (p : a = a') (c : C a b) :
|
definition transportD_eq_transport (p : a = a') (c : C a b) :
|
||||||
p ▹D c = transport (λu, C (u.1) (u.2)) (dpair_eq_dpair p idp) c :=
|
p ▹D c = transport (λu, C (u.1) (u.2)) (dpair_eq_dpair p idp) c :=
|
||||||
by cases p; apply idp
|
by cases p; reflexivity
|
||||||
|
|
||||||
definition sigma_eq_eq_sigma_eq {p1 q1 : a = a'} {p2 : p1 ▹ b = b'} {q2 : q1 ▹ b = b'}
|
definition sigma_eq_eq_sigma_eq {p1 q1 : a = a'} {p2 : p1 ▹ b = b'} {q2 : q1 ▹ b = b'}
|
||||||
(r : p1 = q1) (s : r ▹ p2 = q2) : sigma_eq p1 p2 = sigma_eq q1 q2 :=
|
(r : p1 = q1) (s : r ▹ p2 = q2) : sigma_eq p1 p2 = sigma_eq q1 q2 :=
|
||||||
by cases r; cases s; apply idp
|
by cases r; cases s; reflexivity
|
||||||
|
|
||||||
/- A path between paths in a total space is commonly shown component wise. -/
|
/- A path between paths in a total space is commonly shown component wise. -/
|
||||||
definition sigma_eq2 {p q : u = v} (r : p..1 = q..1) (s : r ▹ p..2 = q..2)
|
definition sigma_eq2 {p q : u = v} (r : p..1 = q..1) (s : r ▹ p..2 = q..2)
|
||||||
|
@ -144,9 +144,9 @@ namespace sigma
|
||||||
revert q r s,
|
revert q r s,
|
||||||
cases p, cases u with u1 u2,
|
cases p, cases u with u1 u2,
|
||||||
intro q r s,
|
intro q r s,
|
||||||
apply concat, rotate 1,
|
transitivity sigma_eq q..1 q..2,
|
||||||
apply sigma_eq_eta,
|
apply sigma_eq_eq_sigma_eq r s,
|
||||||
apply sigma_eq_eq_sigma_eq r s
|
apply sigma_eq_eta,
|
||||||
end
|
end
|
||||||
|
|
||||||
/- In Coq they often have to give u and v explicitly when using the following definition -/
|
/- In Coq they often have to give u and v explicitly when using the following definition -/
|
||||||
|
@ -162,20 +162,19 @@ namespace sigma
|
||||||
|
|
||||||
definition transport_eq (p : a = a') (bc : Σ(b : B a), C a b)
|
definition transport_eq (p : a = a') (bc : Σ(b : B a), C a b)
|
||||||
: p ▹ bc = ⟨p ▹ bc.1, p ▹D bc.2⟩ :=
|
: p ▹ bc = ⟨p ▹ bc.1, p ▹D bc.2⟩ :=
|
||||||
by cases p; cases bc; apply idp
|
by cases p; cases bc; reflexivity
|
||||||
|
|
||||||
/- The special case when the second variable doesn't depend on the first is simpler. -/
|
/- The special case when the second variable doesn't depend on the first is simpler. -/
|
||||||
definition tr_eq_nondep {B : Type} {C : A → B → Type} (p : a = a') (bc : Σ(b : B), C a b)
|
definition tr_eq_nondep {B : Type} {C : A → B → Type} (p : a = a') (bc : Σ(b : B), C a b)
|
||||||
: p ▹ bc = ⟨bc.1, p ▹ bc.2⟩ :=
|
: p ▹ bc = ⟨bc.1, p ▹ bc.2⟩ :=
|
||||||
by cases p; cases bc; apply idp
|
by cases p; cases bc; reflexivity
|
||||||
|
|
||||||
/- Or if the second variable contains a first component that doesn't depend on the first. -/
|
/- Or if the second variable contains a first component that doesn't depend on the first. -/
|
||||||
|
|
||||||
definition tr_eq2_nondep {C : A → Type} {D : Π a:A, B a → C a → Type} (p : a = a')
|
definition tr_eq2_nondep {C : A → Type} {D : Π a:A, B a → C a → Type} (p : a = a')
|
||||||
(bcd : Σ(b : B a) (c : C a), D a b c) : p ▹ bcd = ⟨p ▹ bcd.1, p ▹ bcd.2.1, p ▹D2 bcd.2.2⟩ :=
|
(bcd : Σ(b : B a) (c : C a), D a b c) : p ▹ bcd = ⟨p ▹ bcd.1, p ▹ bcd.2.1, p ▹D2 bcd.2.2⟩ :=
|
||||||
begin
|
begin
|
||||||
cases p, cases bcd with b cd,
|
cases p, cases bcd with b cd, cases cd, reflexivity
|
||||||
cases cd, apply idp
|
|
||||||
end
|
end
|
||||||
|
|
||||||
/- Functorial action -/
|
/- Functorial action -/
|
||||||
|
@ -238,7 +237,7 @@ namespace sigma
|
||||||
: ap (sigma.sigma_functor f g) (sigma_eq p q)
|
: ap (sigma.sigma_functor f g) (sigma_eq p q)
|
||||||
= sigma_eq (ap f p)
|
= sigma_eq (ap f p)
|
||||||
((transport_compose _ f p (g a b))⁻¹ ⬝ (fn_tr_eq_tr_fn p g b)⁻¹ ⬝ ap (g a') q) :=
|
((transport_compose _ f p (g a b))⁻¹ ⬝ (fn_tr_eq_tr_fn p g b)⁻¹ ⬝ ap (g a') q) :=
|
||||||
by cases p; cases q; apply idp
|
by cases p; cases q; reflexivity
|
||||||
|
|
||||||
definition ap_sigma_functor_eq (p : u.1 = v.1) (q : p ▹ u.2 = v.2)
|
definition ap_sigma_functor_eq (p : u.1 = v.1) (q : p ▹ u.2 = v.2)
|
||||||
: ap (sigma.sigma_functor f g) (sigma_eq p q) =
|
: ap (sigma.sigma_functor f g) (sigma_eq p q) =
|
||||||
|
@ -275,8 +274,8 @@ namespace sigma
|
||||||
equiv.mk _ (adjointify
|
equiv.mk _ (adjointify
|
||||||
(λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩)
|
(λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩)
|
||||||
(λuc, ⟨uc.1.1, uc.1.2, !eta⁻¹ ▹ uc.2⟩)
|
(λuc, ⟨uc.1.1, uc.1.2, !eta⁻¹ ▹ uc.2⟩)
|
||||||
begin intro uc; cases uc with u c; cases u; apply idp end
|
begin intro uc, cases uc with u c, cases u, reflexivity end
|
||||||
begin intro av; cases av with a v; cases v; apply idp end)
|
begin intro av, cases av with a v, cases v, reflexivity end)
|
||||||
|
|
||||||
open prod
|
open prod
|
||||||
definition assoc_equiv_prod (C : (A × A') → Type) : (Σa a', C (a,a')) ≃ (Σu, C u) :=
|
definition assoc_equiv_prod (C : (A × A') → Type) : (Σa a', C (a,a')) ≃ (Σu, C u) :=
|
||||||
|
@ -368,7 +367,7 @@ namespace sigma
|
||||||
eapply (trunc_index.rec_on n),
|
eapply (trunc_index.rec_on n),
|
||||||
intro A B HA HB,
|
intro A B HA HB,
|
||||||
fapply is_trunc.is_trunc_equiv_closed,
|
fapply is_trunc.is_trunc_equiv_closed,
|
||||||
apply equiv.symm,
|
symmetry,
|
||||||
apply sigma_equiv_of_is_contr_pr1,
|
apply sigma_equiv_of_is_contr_pr1,
|
||||||
intro n IH A B HA HB,
|
intro n IH A B HA HB,
|
||||||
fapply is_trunc.is_trunc_succ_intro, intro u v,
|
fapply is_trunc.is_trunc_succ_intro, intro u v,
|
||||||
|
|
|
@ -15,22 +15,22 @@ p₁ = p₂ ∨ swap p₁ = p₂
|
||||||
|
|
||||||
infix `~` := eqv -- this is "~"
|
infix `~` := eqv -- this is "~"
|
||||||
|
|
||||||
private theorem eqv.refl {A : Type} : ∀ p : A × A, p ~ p :=
|
private theorem eqv.refl [refl] {A : Type} : ∀ p : A × A, p ~ p :=
|
||||||
take p, or.inl rfl
|
take p, or.inl rfl
|
||||||
|
|
||||||
private theorem eqv.symm {A : Type} : ∀ p₁ p₂ : A × A, p₁ ~ p₂ → p₂ ~ p₁ :=
|
private theorem eqv.symm [symm] {A : Type} : ∀ p₁ p₂ : A × A, p₁ ~ p₂ → p₂ ~ p₁ :=
|
||||||
take p₁ p₂ h, or.elim h
|
take p₁ p₂ h, or.elim h
|
||||||
(λ e, by rewrite e; apply eqv.refl)
|
(λ e, by rewrite e; reflexivity)
|
||||||
(λ e, begin esimp [eqv], rewrite [-e, swap_swap], exact (or.inr rfl) end)
|
(λ e, begin esimp [eqv], rewrite [-e, swap_swap], right, reflexivity end)
|
||||||
|
|
||||||
private theorem eqv.trans {A : Type} : ∀ p₁ p₂ p₃ : A × A, p₁ ~ p₂ → p₂ ~ p₃ → p₁ ~ p₃ :=
|
private theorem eqv.trans [trans] {A : Type} : ∀ p₁ p₂ p₃ : A × A, p₁ ~ p₂ → p₂ ~ p₃ → p₁ ~ p₃ :=
|
||||||
take p₁ p₂ p₃ h₁ h₂, or.elim h₁
|
take p₁ p₂ p₃ h₁ h₂, or.elim h₁
|
||||||
(λ e₁₂, or.elim h₂
|
(λ e₁₂, or.elim h₂
|
||||||
(λ e₂₃, by rewrite [e₁₂, e₂₃]; apply eqv.refl)
|
(λ e₂₃, by rewrite [e₁₂, e₂₃]; reflexivity)
|
||||||
(λ es₂₃, begin esimp [eqv], rewrite -es₂₃, exact (or.inr (congr_arg swap e₁₂)) end))
|
(λ es₂₃, begin esimp [eqv], rewrite -es₂₃, right, congruence, assumption end))
|
||||||
(λ es₁₂, or.elim h₂
|
(λ es₁₂, or.elim h₂
|
||||||
(λ e₂₃, begin esimp [eqv], rewrite es₁₂, exact (or.inr e₂₃) end)
|
(λ e₂₃, begin esimp [eqv], rewrite es₁₂, right, assumption end)
|
||||||
(λ es₂₃, begin esimp [eqv], rewrite [-es₁₂ at es₂₃, swap_swap at es₂₃], exact (or.inl es₂₃) end))
|
(λ es₂₃, begin esimp [eqv], rewrite [-es₁₂ at es₂₃, swap_swap at es₂₃], left, assumption end))
|
||||||
|
|
||||||
private theorem is_equivalence (A : Type) : equivalence (@eqv A) :=
|
private theorem is_equivalence (A : Type) : equivalence (@eqv A) :=
|
||||||
mk_equivalence (@eqv A) (@eqv.refl A) (@eqv.symm A) (@eqv.trans A)
|
mk_equivalence (@eqv A) (@eqv.refl A) (@eqv.symm A) (@eqv.trans A)
|
||||||
|
@ -83,7 +83,7 @@ namespace uprod
|
||||||
| (a₁, b₁) (a₂, b₂) h :=
|
| (a₁, b₁) (a₂, b₂) h :=
|
||||||
or.elim h
|
or.elim h
|
||||||
(λ e, by rewrite e)
|
(λ e, by rewrite e)
|
||||||
(λ es, begin rewrite -es, apply quot.sound, exact (or.inr rfl) end)
|
(λ es, begin rewrite -es, apply quot.sound, right, reflexivity end)
|
||||||
|
|
||||||
definition map {A B : Type} (f : A → B) (u : uprod A) : uprod B :=
|
definition map {A B : Type} (f : A → B) (u : uprod A) : uprod B :=
|
||||||
quot.lift_on u (λ p, map_fn f p) (λ p₁ p₂ c, map_coherent f c)
|
quot.lift_on u (λ p, map_fn f p) (λ p₁ p₂ c, map_coherent f c)
|
||||||
|
@ -95,5 +95,5 @@ namespace uprod
|
||||||
or.inr rfl
|
or.inr rfl
|
||||||
|
|
||||||
theorem map_map {A B C : Type} (g : B → C) (f : A → B) (u : uprod A) : map g (map f u) = map (g ∘ f) u :=
|
theorem map_map {A B C : Type} (g : B → C) (f : A → B) (u : uprod A) : map g (map f u) = map (g ∘ f) u :=
|
||||||
quot.induction_on u (λ p : A × A, begin cases p, apply rfl end)
|
quot.induction_on u (λ p : A × A, begin cases p, reflexivity end)
|
||||||
end uprod
|
end uprod
|
||||||
|
|
Loading…
Reference in a new issue