feat(hit/quotient): prove the flattening lemma
This commit is contained in:
parent
aa9f32a3bd
commit
a99a99f047
10 changed files with 291 additions and 42 deletions
|
@ -20,7 +20,7 @@ The rows indicate the chapters, the columns the sections.
|
||||||
| Ch 3 | + | + | + | + | ½ | + | + | - | + | . | + | | | | |
|
| Ch 3 | + | + | + | + | ½ | + | + | - | + | . | + | | | | |
|
||||||
| Ch 4 | - | + | + | + | . | + | ½ | + | + | | | | | | |
|
| Ch 4 | - | + | + | + | . | + | ½ | + | + | | | | | | |
|
||||||
| Ch 5 | - | . | ½ | - | - | . | . | ½ | | | | | | | |
|
| Ch 5 | - | . | ½ | - | - | . | . | ½ | | | | | | | |
|
||||||
| Ch 6 | . | + | + | + | + | ½ | ½ | + | ¾ | ¼ | ¾ | - | . | | |
|
| Ch 6 | . | + | + | + | + | ½ | ½ | + | ¾ | ¼ | ¾ | + | . | | |
|
||||||
| Ch 7 | + | + | + | - | ½ | - | - | | | | | | | | |
|
| Ch 7 | + | + | + | - | ½ | - | - | | | | | | | | |
|
||||||
| Ch 8 | ¾ | - | - | - | - | - | - | - | - | - | | | | | |
|
| Ch 8 | ¾ | - | - | - | - | - | - | - | - | - | | | | | |
|
||||||
| Ch 9 | ¾ | + | + | ½ | ¾ | ½ | - | - | - | | | | | | |
|
| Ch 9 | ¾ | + | + | ½ | ¾ | ½ | - | - | - | | | | | | |
|
||||||
|
@ -120,7 +120,7 @@ Chapter 6: Higher inductive types
|
||||||
- 6.9 (Truncations): [hit.trunc](hit/trunc.hlean) (except Lemma 6.9.3)
|
- 6.9 (Truncations): [hit.trunc](hit/trunc.hlean) (except Lemma 6.9.3)
|
||||||
- 6.10 (Quotients): [hit.set_quotient](hit/set_quotient.hlean) (up to 6.10.3). We define integers differently, to make them compute, in the folder [types.int](types/int/int.md). 6.10.13 is in [types.int.hott](types/int/hott.hlean)
|
- 6.10 (Quotients): [hit.set_quotient](hit/set_quotient.hlean) (up to 6.10.3). We define integers differently, to make them compute, in the folder [types.int](types/int/int.md). 6.10.13 is in [types.int.hott](types/int/hott.hlean)
|
||||||
- 6.11 (Algebra): [algebra.group](algebra/group.hlean), [algebra.fundamental_group](algebra/fundamental_group.hlean) (no homotopy groups yet)
|
- 6.11 (Algebra): [algebra.group](algebra/group.hlean), [algebra.fundamental_group](algebra/fundamental_group.hlean) (no homotopy groups yet)
|
||||||
- 6.12 (The flattening lemma): not formalized yet
|
- 6.12 (The flattening lemma): [hit.quotient](hit/quotient.hlean) (for quotients instead of homotopy coequalizers, but these are practically the same)
|
||||||
- 6.13 (The general syntax of higher inductive definitions): no formalizable content
|
- 6.13 (The general syntax of higher inductive definitions): no formalizable content
|
||||||
|
|
||||||
Chapter 7: Homotopy n-types
|
Chapter 7: Homotopy n-types
|
||||||
|
|
45
hott/cubical/pathover2.hlean
Normal file
45
hott/cubical/pathover2.hlean
Normal file
|
@ -0,0 +1,45 @@
|
||||||
|
/-
|
||||||
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
||||||
|
Released under Apache 2.0 license as described in the file LICENSE.
|
||||||
|
Author: Floris van Doorn
|
||||||
|
|
||||||
|
Coherence conditions for operations on pathovers
|
||||||
|
-/
|
||||||
|
|
||||||
|
open function
|
||||||
|
|
||||||
|
namespace eq
|
||||||
|
|
||||||
|
variables {A A' A'' : Type} {a a' a₂ : A} (B : A → Type) {p : a = a₂}
|
||||||
|
{b b' : B a} {b₂ : B a₂}
|
||||||
|
|
||||||
|
definition pathover_ap_id (q : b =[p] b₂) : pathover_ap B id q = change_path (ap_id p)⁻¹ q :=
|
||||||
|
by induction q; reflexivity
|
||||||
|
|
||||||
|
definition pathover_ap_compose (B : A'' → Type) (g : A' → A'') (f : A → A')
|
||||||
|
{b : B (g (f a))} {b₂ : B (g (f a₂))} (q : b =[p] b₂) : pathover_ap B (g ∘ f) q
|
||||||
|
= change_path (ap_compose g f p)⁻¹ (pathover_ap B g (pathover_ap (B ∘ g) f q)) :=
|
||||||
|
by induction q; reflexivity
|
||||||
|
|
||||||
|
definition pathover_ap_compose_rev (B : A'' → Type) (g : A' → A'') (f : A → A')
|
||||||
|
{b : B (g (f a))} {b₂ : B (g (f a₂))} (q : b =[p] b₂) :
|
||||||
|
pathover_ap B g (pathover_ap (B ∘ g) f q)
|
||||||
|
= change_path (ap_compose g f p) (pathover_ap B (g ∘ f) q) :=
|
||||||
|
by induction q; reflexivity
|
||||||
|
|
||||||
|
definition apdo_eq_apdo_ap {B : A' → Type} (g : Πb, B b) (f : A → A') (p : a = a')
|
||||||
|
: apdo (λx, g (f x)) p = pathover_of_pathover_ap B f (apdo g (ap f p)) :=
|
||||||
|
by induction p; reflexivity
|
||||||
|
|
||||||
|
definition pathover_of_tr_eq_idp (r : b = b') : pathover_of_tr_eq r = pathover_idp_of_eq r :=
|
||||||
|
idp
|
||||||
|
|
||||||
|
definition pathover_of_tr_eq_eq_concato (r : p ▸ b = b₂)
|
||||||
|
: pathover_of_tr_eq r = pathover_tr p b ⬝o pathover_idp_of_eq r :=
|
||||||
|
by induction r; induction p; reflexivity
|
||||||
|
|
||||||
|
definition apo011_eq_apo11_apdo (f : Πa, B a → A') (p : a = a₂) (q : b =[p] b₂)
|
||||||
|
: apo011 f p q = eq_of_pathover (apo11 (apdo f p) q) :=
|
||||||
|
by induction q; reflexivity
|
||||||
|
|
||||||
|
end eq
|
|
@ -280,9 +280,13 @@ namespace eq
|
||||||
{ reflexivity}
|
{ reflexivity}
|
||||||
end
|
end
|
||||||
|
|
||||||
-- example (p q : a = a') : to_inv (hdeg_square_equiv' p q) = hdeg_square := idp -- this fails
|
|
||||||
example (p q : a = a') : to_inv (hdeg_square_equiv p q) = hdeg_square := idp
|
example (p q : a = a') : to_inv (hdeg_square_equiv p q) = hdeg_square := idp
|
||||||
|
|
||||||
|
/-
|
||||||
|
characterization of pathovers in a equality type. The type B of the equality is fixed here.
|
||||||
|
A version where B may also varies over the path p is given in the file squareover
|
||||||
|
-/
|
||||||
|
|
||||||
definition eq_pathover [unfold 7] {f g : A → B} {p : a = a'} {q : f a = g a} {r : f a' = g a'}
|
definition eq_pathover [unfold 7] {f g : A → B} {p : a = a'} {q : f a = g a} {r : f a' = g a'}
|
||||||
(s : square q r (ap f p) (ap g p)) : q =[p] r :=
|
(s : square q r (ap f p) (ap g p)) : q =[p] r :=
|
||||||
by induction p;apply pathover_idp_of_eq;exact eq_of_vdeg_square s
|
by induction p;apply pathover_idp_of_eq;exact eq_of_vdeg_square s
|
||||||
|
|
|
@ -140,7 +140,7 @@ namespace eq
|
||||||
-/
|
-/
|
||||||
|
|
||||||
definition square_of_squareover_ids {b₀₀ b₀₂ b₂₀ b₂₂ : B a}
|
definition square_of_squareover_ids {b₀₀ b₀₂ b₂₀ b₂₂ : B a}
|
||||||
(t : b₀₀ = b₂₀) (b : b₀₂ = b₂₂) (l : b₀₀ = b₀₂) (r : b₂₀ = b₂₂)
|
{t : b₀₀ = b₂₀} {b : b₀₂ = b₂₂} {l : b₀₀ = b₀₂} {r : b₂₀ = b₂₂}
|
||||||
(so : squareover B ids (pathover_idp_of_eq t)
|
(so : squareover B ids (pathover_idp_of_eq t)
|
||||||
(pathover_idp_of_eq b)
|
(pathover_idp_of_eq b)
|
||||||
(pathover_idp_of_eq l)
|
(pathover_idp_of_eq l)
|
||||||
|
@ -153,7 +153,7 @@ namespace eq
|
||||||
end
|
end
|
||||||
|
|
||||||
definition squareover_ids_of_square {b₀₀ b₀₂ b₂₀ b₂₂ : B a}
|
definition squareover_ids_of_square {b₀₀ b₀₂ b₂₀ b₂₂ : B a}
|
||||||
(t : b₀₀ = b₂₀) (b : b₀₂ = b₂₂) (l : b₀₀ = b₀₂) (r : b₂₀ = b₂₂) (q : square t b l r)
|
{t : b₀₀ = b₂₀} {b : b₀₂ = b₂₂} {l : b₀₀ = b₀₂} {r : b₂₀ = b₂₂} (q : square t b l r)
|
||||||
: squareover B ids (pathover_idp_of_eq t)
|
: squareover B ids (pathover_idp_of_eq t)
|
||||||
(pathover_idp_of_eq b)
|
(pathover_idp_of_eq b)
|
||||||
(pathover_idp_of_eq l)
|
(pathover_idp_of_eq l)
|
||||||
|
@ -241,8 +241,16 @@ namespace eq
|
||||||
-- : squareover B s₁₁ q₁₀ q₁₂ !pathover_tr !pathover_tr :=
|
-- : squareover B s₁₁ q₁₀ q₁₂ !pathover_tr !pathover_tr :=
|
||||||
-- by induction p;exact vrflo
|
-- by induction p;exact vrflo
|
||||||
|
|
||||||
|
/- A version of eq_pathover where the type of the equality also varies -/
|
||||||
|
definition eq_pathover_dep {f g : Πa, B a} {p : a = a'} {q : f a = g a}
|
||||||
|
{r : f a' = g a'} (s : squareover B hrfl (pathover_idp_of_eq q) (pathover_idp_of_eq r)
|
||||||
|
(apdo f p) (apdo g p)) : q =[p] r :=
|
||||||
|
begin
|
||||||
|
induction p, apply pathover_idp_of_eq, apply eq_of_vdeg_square, exact square_of_squareover_ids s
|
||||||
|
end
|
||||||
|
|
||||||
/- charcaterization of pathovers in pathovers -/
|
/- charcaterization of pathovers in pathovers -/
|
||||||
-- in this version the fibration (B) of the pathover does not depend on the variable a
|
-- in this version the fibration (B) of the pathover does not depend on the variable (a)
|
||||||
definition pathover_pathover {a' a₂' : A'} {p : a' = a₂'} {f g : A' → A}
|
definition pathover_pathover {a' a₂' : A'} {p : a' = a₂'} {f g : A' → A}
|
||||||
{b : Πa, B (f a)} {b₂ : Πa, B (g a)} {q : Π(a' : A'), f a' = g a'}
|
{b : Πa, B (f a)} {b₂ : Πa, B (g a)} {q : Π(a' : A'), f a' = g a'}
|
||||||
(r : pathover B (b a') (q a') (b₂ a'))
|
(r : pathover B (b a') (q a') (b₂ a'))
|
||||||
|
|
|
@ -14,7 +14,9 @@ See also .set_quotient
|
||||||
* eq_of_rel : Π{a a' : A}, R a a' → class_of a = class_of a' (R explicit)
|
* eq_of_rel : Π{a a' : A}, R a a' → class_of a = class_of a' (R explicit)
|
||||||
-/
|
-/
|
||||||
|
|
||||||
open eq equiv sigma.ops
|
import arity cubical.squareover types.arrow cubical.pathover2
|
||||||
|
|
||||||
|
open eq equiv sigma sigma.ops equiv.ops pi
|
||||||
|
|
||||||
namespace quotient
|
namespace quotient
|
||||||
|
|
||||||
|
@ -44,11 +46,21 @@ namespace quotient
|
||||||
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') : Type :=
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') : Type :=
|
||||||
quotient.elim_type Pc Pp x
|
quotient.elim_type Pc Pp x
|
||||||
|
|
||||||
theorem elim_type_eq_of_rel (Pc : A → Type)
|
theorem elim_type_eq_of_rel_fn (Pc : A → Type)
|
||||||
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') {a a' : A} (H : R a a')
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') {a a' : A} (H : R a a')
|
||||||
: transport (quotient.elim_type Pc Pp) (eq_of_rel R H) = to_fun (Pp H) :=
|
: transport (quotient.elim_type Pc Pp) (eq_of_rel R H) = to_fun (Pp H) :=
|
||||||
by rewrite [tr_eq_cast_ap_fn, ↑quotient.elim_type, elim_eq_of_rel];apply cast_ua_fn
|
by rewrite [tr_eq_cast_ap_fn, ↑quotient.elim_type, elim_eq_of_rel];apply cast_ua_fn
|
||||||
|
|
||||||
|
theorem elim_type_eq_of_rel.{u} (Pc : A → Type.{u})
|
||||||
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') {a a' : A} (H : R a a') (p : Pc a)
|
||||||
|
: transport (quotient.elim_type Pc Pp) (eq_of_rel R H) p = to_fun (Pp H) p :=
|
||||||
|
ap10 (elim_type_eq_of_rel_fn Pc Pp H) p
|
||||||
|
|
||||||
|
definition elim_type_eq_of_rel' (Pc : A → Type)
|
||||||
|
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a') {a a' : A} (H : R a a') (p : Pc a)
|
||||||
|
: pathover (quotient.elim_type Pc Pp) p (eq_of_rel R H) (to_fun (Pp H) p) :=
|
||||||
|
pathover_of_tr_eq (elim_type_eq_of_rel Pc Pp H p)
|
||||||
|
|
||||||
definition elim_type_uncurried (H : Σ(Pc : A → Type), Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a')
|
definition elim_type_uncurried (H : Σ(Pc : A → Type), Π⦃a a' : A⦄ (H : R a a'), Pc a ≃ Pc a')
|
||||||
: quotient R → Type :=
|
: quotient R → Type :=
|
||||||
quotient.elim_type H.1 H.2
|
quotient.elim_type H.1 H.2
|
||||||
|
@ -59,3 +71,120 @@ attribute quotient.elim [unfold 6] [recursor 6]
|
||||||
attribute quotient.elim_type [unfold 5]
|
attribute quotient.elim_type [unfold 5]
|
||||||
attribute quotient.elim_on [unfold 4]
|
attribute quotient.elim_on [unfold 4]
|
||||||
attribute quotient.elim_type_on [unfold 3]
|
attribute quotient.elim_type_on [unfold 3]
|
||||||
|
|
||||||
|
namespace quotient
|
||||||
|
|
||||||
|
section
|
||||||
|
variables {A : Type} (R : A → A → Type)
|
||||||
|
|
||||||
|
/- The dependent universal property -/
|
||||||
|
definition quotient_pi_equiv (C : quotient R → Type) : (Πx, C x) ≃
|
||||||
|
(Σ(f : Π(a : A), C (class_of R a)), Π⦃a a' : A⦄ (H : R a a'), f a =[eq_of_rel R H] f a') :=
|
||||||
|
begin
|
||||||
|
fapply equiv.MK,
|
||||||
|
{ intro f, exact ⟨λa, f (class_of R a), λa a' H, apdo f (eq_of_rel R H)⟩},
|
||||||
|
{ intro v x, induction v with i p, induction x,
|
||||||
|
exact (i a),
|
||||||
|
exact (p H)},
|
||||||
|
{ intro v, induction v with i p, esimp,
|
||||||
|
apply ap (sigma.mk i), apply eq_of_homotopy3, intro a a' H, apply rec_eq_of_rel},
|
||||||
|
{ intro f, apply eq_of_homotopy, intro x, induction x: esimp,
|
||||||
|
apply eq_pathover_dep, esimp, rewrite rec_eq_of_rel, exact hrflo},
|
||||||
|
end
|
||||||
|
end
|
||||||
|
|
||||||
|
/- the flattening lemma -/
|
||||||
|
|
||||||
|
namespace flattening
|
||||||
|
section
|
||||||
|
|
||||||
|
parameters {A : Type} (R : A → A → Type) (C : A → Type) (f : Π⦃a a'⦄, R a a' → C a ≃ C a')
|
||||||
|
include f
|
||||||
|
variables {a a' : A} {r : R a a'}
|
||||||
|
|
||||||
|
local abbreviation P [unfold 5] := quotient.elim_type C f
|
||||||
|
|
||||||
|
definition flattening_type : Type := Σa, C a
|
||||||
|
local abbreviation X := flattening_type
|
||||||
|
inductive flattening_rel : X → X → Type :=
|
||||||
|
| mk : Π⦃a a' : A⦄ (r : R a a') (c : C a), flattening_rel ⟨a, c⟩ ⟨a', f r c⟩
|
||||||
|
|
||||||
|
definition Ppt [constructor] (c : C a) : sigma P :=
|
||||||
|
⟨class_of R a, c⟩
|
||||||
|
|
||||||
|
definition Peq (r : R a a') (c : C a) : Ppt c = Ppt (f r c) :=
|
||||||
|
begin
|
||||||
|
fapply sigma_eq: esimp,
|
||||||
|
{ apply eq_of_rel R r},
|
||||||
|
{ refine elim_type_eq_of_rel' C f r c}
|
||||||
|
end
|
||||||
|
|
||||||
|
definition rec {Q : sigma P → Type} (Qpt : Π{a : A} (x : C a), Q (Ppt x))
|
||||||
|
(Qeq : Π⦃a a' : A⦄ (r : R a a') (c : C a), Qpt c =[Peq r c] Qpt (f r c))
|
||||||
|
(v : sigma P) : Q v :=
|
||||||
|
begin
|
||||||
|
induction v with q p,
|
||||||
|
induction q,
|
||||||
|
{ exact Qpt p},
|
||||||
|
{ apply pi_pathover_left', esimp, intro c,
|
||||||
|
refine _ ⬝op apd Qpt (elim_type_eq_of_rel C f H c)⁻¹ᵖ,
|
||||||
|
refine _ ⬝op (tr_compose Q Ppt _ _)⁻¹ ,
|
||||||
|
rewrite ap_inv,
|
||||||
|
refine pathover_cancel_right _ !tr_pathover⁻¹ᵒ,
|
||||||
|
refine change_path _ (Qeq H c),
|
||||||
|
symmetry, rewrite [↑[Ppt, Peq]],
|
||||||
|
refine whisker_left _ !ap_dpair ⬝ _,
|
||||||
|
refine !dpair_eq_dpair_con⁻¹ ⬝ _, esimp,
|
||||||
|
apply ap (dpair_eq_dpair _),
|
||||||
|
esimp [elim_type_eq_of_rel',pathover_idp_of_eq],
|
||||||
|
exact !pathover_of_tr_eq_eq_concato⁻¹},
|
||||||
|
end
|
||||||
|
|
||||||
|
definition elim {Q : Type} (Qpt : Π{a : A}, C a → Q)
|
||||||
|
(Qeq : Π⦃a a' : A⦄ (r : R a a') (c : C a), Qpt c = Qpt (f r c))
|
||||||
|
(v : sigma P) : Q :=
|
||||||
|
begin
|
||||||
|
induction v with q p,
|
||||||
|
induction q,
|
||||||
|
{ exact Qpt p},
|
||||||
|
{ apply arrow_pathover_constant_right, esimp,
|
||||||
|
intro c, exact Qeq H c ⬝ ap Qpt (elim_type_eq_of_rel C f H c)⁻¹},
|
||||||
|
end
|
||||||
|
|
||||||
|
theorem elim_Peq {Q : Type} (Qpt : Π{a : A}, C a → Q)
|
||||||
|
(Qeq : Π⦃a a' : A⦄ (r : R a a') (c : C a), Qpt c = Qpt (f r c)) {a a' : A} (r : R a a')
|
||||||
|
(c : C a) : ap (elim @Qpt Qeq) (Peq r c) = Qeq r c :=
|
||||||
|
begin
|
||||||
|
refine !ap_dpair_eq_dpair ⬝ _,
|
||||||
|
rewrite [apo011_eq_apo11_apdo, rec_eq_of_rel, ▸*, apo011_arrow_pathover_constant_right,
|
||||||
|
↑elim_type_eq_of_rel', to_right_inv !pathover_equiv_tr_eq, ap_inv],
|
||||||
|
apply inv_con_cancel_right
|
||||||
|
end
|
||||||
|
|
||||||
|
open flattening_rel
|
||||||
|
definition flattening_lemma : sigma P ≃ quotient flattening_rel :=
|
||||||
|
begin
|
||||||
|
fapply equiv.MK,
|
||||||
|
{ refine elim _ _,
|
||||||
|
{ intro a c, exact class_of _ ⟨a, c⟩},
|
||||||
|
{ intro a a' r c, apply eq_of_rel, constructor}},
|
||||||
|
{ intro q, induction q with x x x' H,
|
||||||
|
{ exact Ppt x.2},
|
||||||
|
{ induction H, esimp, apply Peq}},
|
||||||
|
{ intro q, induction q with x x x' H: esimp,
|
||||||
|
{ induction x with a c, reflexivity},
|
||||||
|
{ induction H, esimp, apply eq_pathover, apply hdeg_square,
|
||||||
|
refine ap_compose (elim _ _) (quotient.elim _ _) _ ⬝ _,
|
||||||
|
rewrite [elim_eq_of_rel, ap_id, ▸*],
|
||||||
|
apply elim_Peq}},
|
||||||
|
{ refine rec (λa x, idp) _, intros,
|
||||||
|
apply eq_pathover, apply hdeg_square,
|
||||||
|
refine ap_compose (quotient.elim _ _) (elim _ _) _ ⬝ _,
|
||||||
|
rewrite [elim_Peq, ap_id, ▸*],
|
||||||
|
apply elim_eq_of_rel}
|
||||||
|
end
|
||||||
|
|
||||||
|
end
|
||||||
|
end flattening
|
||||||
|
|
||||||
|
end quotient
|
||||||
|
|
|
@ -12,7 +12,7 @@ import .path .equiv
|
||||||
open equiv is_equiv equiv.ops function
|
open equiv is_equiv equiv.ops function
|
||||||
|
|
||||||
variables {A A' : Type} {B B' : A → Type} {B'' : A' → Type} {C : Π⦃a⦄, B a → Type}
|
variables {A A' : Type} {B B' : A → Type} {B'' : A' → Type} {C : Π⦃a⦄, B a → Type}
|
||||||
{a a₂ a₃ a₄ : A} {p p' : a = a₂} {p₂ : a₂ = a₃} {p₃ : a₃ = a₄}
|
{a a₂ a₃ a₄ : A} {p p' : a = a₂} {p₂ : a₂ = a₃} {p₃ : a₃ = a₄} {p₁₃ : a = a₃}
|
||||||
{b b' : B a} {b₂ b₂' : B a₂} {b₃ : B a₃} {b₄ : B a₄}
|
{b b' : B a} {b₂ b₂' : B a₂} {b₃ : B a₃} {b₄ : B a₄}
|
||||||
{c : C b} {c₂ : C b₂}
|
{c : C b} {c₂ : C b₂}
|
||||||
|
|
||||||
|
@ -79,13 +79,28 @@ namespace eq
|
||||||
definition eq_concato [unfold 9] (q : b = b') (r : b' =[p] b₂) : b =[p] b₂ :=
|
definition eq_concato [unfold 9] (q : b = b') (r : b' =[p] b₂) : b =[p] b₂ :=
|
||||||
by induction q;exact r
|
by induction q;exact r
|
||||||
|
|
||||||
|
definition change_path [unfold 9] (q : p = p') (r : b =[p] b₂) : b =[p'] b₂ :=
|
||||||
|
q ▸ r
|
||||||
|
|
||||||
-- infix ` ⬝ ` := concato
|
-- infix ` ⬝ ` := concato
|
||||||
infix `⬝o`:75 := concato
|
infix ` ⬝o `:72 := concato
|
||||||
infix `⬝op`:75 := concato_eq
|
infix ` ⬝op `:73 := concato_eq
|
||||||
infix `⬝po`:75 := eq_concato
|
infix ` ⬝po `:73 := eq_concato
|
||||||
-- postfix `⁻¹` := inverseo
|
-- postfix `⁻¹` := inverseo
|
||||||
postfix `⁻¹ᵒ`:(max+10) := inverseo
|
postfix `⁻¹ᵒ`:(max+10) := inverseo
|
||||||
|
|
||||||
|
definition pathover_cancel_right (q : b =[p ⬝ p₂] b₃) (r : b₃ =[p₂⁻¹] b₂) : b =[p] b₂ :=
|
||||||
|
change_path !con_inv_cancel_right (q ⬝o r)
|
||||||
|
|
||||||
|
definition pathover_cancel_right' (q : b =[p₁₃ ⬝ p₂⁻¹] b₂) (r : b₂ =[p₂] b₃) : b =[p₁₃] b₃ :=
|
||||||
|
change_path !inv_con_cancel_right (q ⬝o r)
|
||||||
|
|
||||||
|
definition pathover_cancel_left (q : b₂ =[p⁻¹] b) (r : b =[p ⬝ p₂] b₃) : b₂ =[p₂] b₃ :=
|
||||||
|
change_path !inv_con_cancel_left (q ⬝o r)
|
||||||
|
|
||||||
|
definition pathover_cancel_left' (q : b =[p] b₂) (r : b₂ =[p⁻¹ ⬝ p₁₃] b₃) : b =[p₁₃] b₃ :=
|
||||||
|
change_path !con_inv_cancel_left (q ⬝o r)
|
||||||
|
|
||||||
/- Some of the theorems analogous to theorems for = in init.path -/
|
/- Some of the theorems analogous to theorems for = in init.path -/
|
||||||
|
|
||||||
definition cono_idpo (r : b =[p] b₂) : r ⬝o idpo =[con_idp p] r :=
|
definition cono_idpo (r : b =[p] b₂) : r ⬝o idpo =[con_idp p] r :=
|
||||||
|
@ -111,7 +126,7 @@ namespace eq
|
||||||
definition eq_of_pathover {a' a₂' : A'} (q : a' =[p] a₂') : a' = a₂' :=
|
definition eq_of_pathover {a' a₂' : A'} (q : a' =[p] a₂') : a' = a₂' :=
|
||||||
by cases q;reflexivity
|
by cases q;reflexivity
|
||||||
|
|
||||||
definition pathover_of_eq {a' a₂' : A'} (q : a' = a₂') : a' =[p] a₂' :=
|
definition pathover_of_eq [unfold 5 8] {a' a₂' : A'} (q : a' = a₂') : a' =[p] a₂' :=
|
||||||
by cases p;cases q;constructor
|
by cases p;cases q;constructor
|
||||||
|
|
||||||
definition pathover_constant [constructor] (p : a = a₂) (a' a₂' : A') : a' =[p] a₂' ≃ a' = a₂' :=
|
definition pathover_constant [constructor] (p : a = a₂) (a' a₂' : A') : a' =[p] a₂' ≃ a' = a₂' :=
|
||||||
|
@ -191,10 +206,16 @@ namespace eq
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
definition pathover_of_pathover_tr (q : b =[p ⬝ p₂] p₂ ▸ b₂) : b =[p] b₂ :=
|
definition pathover_of_pathover_tr (q : b =[p ⬝ p₂] p₂ ▸ b₂) : b =[p] b₂ :=
|
||||||
by cases p₂;exact q
|
pathover_cancel_right q !pathover_tr⁻¹ᵒ
|
||||||
|
|
||||||
definition pathover_tr_of_pathover {p : a = a₃} (q : b =[p ⬝ p₂⁻¹] b₂) : b =[p] p₂ ▸ b₂ :=
|
definition pathover_tr_of_pathover (q : b =[p₁₃ ⬝ p₂⁻¹] b₂) : b =[p₁₃] p₂ ▸ b₂ :=
|
||||||
by cases p₂;exact q
|
pathover_cancel_right' q !pathover_tr
|
||||||
|
|
||||||
|
definition pathover_of_tr_pathover (q : p ▸ b =[p⁻¹ ⬝ p₁₃] b₃) : b =[p₁₃] b₃ :=
|
||||||
|
pathover_cancel_left' !pathover_tr q
|
||||||
|
|
||||||
|
definition tr_pathover_of_pathover (q : b =[p ⬝ p₂] b₃) : p ▸ b =[p₂] b₃ :=
|
||||||
|
pathover_cancel_left !pathover_tr⁻¹ᵒ q
|
||||||
|
|
||||||
definition pathover_tr_of_eq (q : b = b') : b =[p] p ▸ b' :=
|
definition pathover_tr_of_eq (q : b = b') : b =[p] p ▸ b' :=
|
||||||
by cases q;apply pathover_tr
|
by cases q;apply pathover_tr
|
||||||
|
@ -216,7 +237,7 @@ namespace eq
|
||||||
(q : b =[p] b₂) : g a b =[p] g a₂ b₂ :=
|
(q : b =[p] b₂) : g a b =[p] g a₂ b₂ :=
|
||||||
by induction q; constructor
|
by induction q; constructor
|
||||||
|
|
||||||
definition apo011 (f : Πa, B a → A') (Ha : a = a₂) (Hb : b =[Ha] b₂)
|
definition apo011 [unfold 10] (f : Πa, B a → A') (Ha : a = a₂) (Hb : b =[Ha] b₂)
|
||||||
: f a b = f a₂ b₂ :=
|
: f a b = f a₂ b₂ :=
|
||||||
by cases Hb; reflexivity
|
by cases Hb; reflexivity
|
||||||
|
|
||||||
|
@ -224,7 +245,7 @@ namespace eq
|
||||||
(Hc : c =[apo011 C Ha Hb] c₂) : f a b c = f a₂ b₂ c₂ :=
|
(Hc : c =[apo011 C Ha Hb] c₂) : f a b c = f a₂ b₂ c₂ :=
|
||||||
by cases Hb; apply (idp_rec_on Hc); apply idp
|
by cases Hb; apply (idp_rec_on Hc); apply idp
|
||||||
|
|
||||||
definition apo11 {f : Πb, C b} {g : Πb₂, C b₂} (r : f =[p] g)
|
definition apod11 {f : Πb, C b} {g : Πb₂, C b₂} (r : f =[p] g)
|
||||||
{b : B a} {b₂ : B a₂} (q : b =[p] b₂) : f b =[apo011 C p q] g b₂ :=
|
{b : B a} {b₂ : B a₂} (q : b =[p] b₂) : f b =[apo011 C p q] g b₂ :=
|
||||||
by cases r; apply (idp_rec_on q); constructor
|
by cases r; apply (idp_rec_on q); constructor
|
||||||
|
|
||||||
|
@ -236,6 +257,10 @@ namespace eq
|
||||||
(b : B a) : f b =[p] g (p ▸ b) :=
|
(b : B a) : f b =[p] g (p ▸ b) :=
|
||||||
by cases r; constructor
|
by cases r; constructor
|
||||||
|
|
||||||
|
definition apo11 {f : B a → B' a} {g : B a₂ → B' a₂} (r : f =[p] g)
|
||||||
|
(q : b =[p] b₂) : f b =[p] g b₂ :=
|
||||||
|
by induction q; exact apo10 r b
|
||||||
|
|
||||||
definition apdo_compose1 (g : Πa, B a → B' a) (f : Πa, B a) (p : a = a₂)
|
definition apdo_compose1 (g : Πa, B a → B' a) (f : Πa, B a) (p : a = a₂)
|
||||||
: apdo (g ∘' f) p = apo g (apdo f p) :=
|
: apdo (g ∘' f) p = apo g (apdo f p) :=
|
||||||
by induction p; reflexivity
|
by induction p; reflexivity
|
||||||
|
@ -260,9 +285,6 @@ namespace eq
|
||||||
: eq_concato q⁻¹ (pathover_idp_of_eq q) = (idpo : b' =[refl a] b') :=
|
: eq_concato q⁻¹ (pathover_idp_of_eq q) = (idpo : b' =[refl a] b') :=
|
||||||
by induction q;constructor
|
by induction q;constructor
|
||||||
|
|
||||||
definition change_path [unfold 9] (q : p = p') (r : b =[p] b₂) : b =[p'] b₂ :=
|
|
||||||
q ▸ r
|
|
||||||
|
|
||||||
definition change_path_equiv (f : Π{a}, B a ≃ B' a) (r : b =[p] b₂) : f b =[p] f b₂ :=
|
definition change_path_equiv (f : Π{a}, B a ≃ B' a) (r : b =[p] b₂) : f b =[p] f b₂ :=
|
||||||
by induction r;constructor
|
by induction r;constructor
|
||||||
|
|
||||||
|
|
|
@ -33,7 +33,7 @@ end unit
|
||||||
-- ----------
|
-- ----------
|
||||||
|
|
||||||
notation `Σ` binders `, ` r:(scoped P, sigma P) := r
|
notation `Σ` binders `, ` r:(scoped P, sigma P) := r
|
||||||
|
abbreviation dpair [constructor] := @sigma.mk
|
||||||
namespace sigma
|
namespace sigma
|
||||||
notation `⟨`:max t:(foldr `, ` (e r, mk e r)) `⟩`:0 := t --input ⟨ ⟩ as \< \>
|
notation `⟨`:max t:(foldr `, ` (e r, mk e r)) `⟩`:0 := t --input ⟨ ⟩ as \< \>
|
||||||
|
|
||||||
|
|
|
@ -13,8 +13,8 @@ open eq equiv is_equiv funext pi equiv.ops is_trunc unit
|
||||||
|
|
||||||
namespace pi
|
namespace pi
|
||||||
|
|
||||||
variables {A A' : Type} {B B' : Type} {C : A → B → Type}
|
variables {A A' : Type} {B B' : Type} {C : A → B → Type} {D : A → Type}
|
||||||
{a a' a'' : A} {b b' b'' : B} {f g : A → B}
|
{a a' a'' : A} {b b' b'' : B} {f g : A → B} {d : D a} {d' : D a'}
|
||||||
|
|
||||||
-- all lemmas here are special cases of the ones for pi-types
|
-- all lemmas here are special cases of the ones for pi-types
|
||||||
|
|
||||||
|
@ -94,7 +94,7 @@ namespace pi
|
||||||
definition arrow_pathover_left {B C : A → Type} {f : B a → C a} {g : B a' → C a'} {p : a = a'}
|
definition arrow_pathover_left {B C : A → Type} {f : B a → C a} {g : B a' → C a'} {p : a = a'}
|
||||||
(r : Π(b : B a), f b =[p] g (p ▸ b)) : f =[p] g :=
|
(r : Π(b : B a), f b =[p] g (p ▸ b)) : f =[p] g :=
|
||||||
begin
|
begin
|
||||||
cases p, apply pathover_idp_of_eq,
|
induction p, apply pathover_idp_of_eq,
|
||||||
apply eq_of_homotopy, intro b,
|
apply eq_of_homotopy, intro b,
|
||||||
exact eq_of_pathover_idp (r b),
|
exact eq_of_pathover_idp (r b),
|
||||||
end
|
end
|
||||||
|
@ -111,10 +111,28 @@ namespace pi
|
||||||
{p : a = a'} (r : Π(b : B), f b =[p] g b) : f =[p] g :=
|
{p : a = a'} (r : Π(b : B), f b =[p] g b) : f =[p] g :=
|
||||||
pi_pathover_constant r
|
pi_pathover_constant r
|
||||||
|
|
||||||
definition arrow_pathover_constant_right {B : A → Type} {C : Type} {f : B a → C} {g : B a' → C}
|
definition arrow_pathover_constant_right' {B : A → Type} {C : Type}
|
||||||
{p : a = a'} (r : Π(b : B a), f b = g (p ▸ b)) : f =[p] g :=
|
{f : B a → C} {g : B a' → C} {p : a = a'}
|
||||||
|
(r : Π⦃b : B a⦄ ⦃b' : B a'⦄ (q : b =[p] b'), f b = g b') : f =[p] g :=
|
||||||
|
arrow_pathover (λb b' q, pathover_of_eq (r q))
|
||||||
|
|
||||||
|
definition arrow_pathover_constant_right {B : A → Type} {C : Type} {f : B a → C}
|
||||||
|
{g : B a' → C} {p : a = a'} (r : Π(b : B a), f b = g (p ▸ b)) : f =[p] g :=
|
||||||
arrow_pathover_left (λb, pathover_of_eq (r b))
|
arrow_pathover_left (λb, pathover_of_eq (r b))
|
||||||
|
|
||||||
|
/- a lemma used for the flattening lemma -/
|
||||||
|
definition apo011_arrow_pathover_constant_right {f : D a → A'} {g : D a' → A'} {p : a = a'}
|
||||||
|
{q : d =[p] d'} (r : Π(d : D a), f d = g (p ▸ d))
|
||||||
|
: eq_of_pathover (apo11 (arrow_pathover_constant_right r) q) = r d ⬝ ap g (tr_eq_of_pathover q)
|
||||||
|
:=
|
||||||
|
begin
|
||||||
|
induction q, esimp at r,
|
||||||
|
eapply homotopy.rec_on r, clear r, esimp, intro r, induction r, esimp,
|
||||||
|
esimp [arrow_pathover_constant_right, arrow_pathover_left],
|
||||||
|
rewrite [eq_of_homotopy_idp]
|
||||||
|
end
|
||||||
|
|
||||||
|
|
||||||
/-
|
/-
|
||||||
The fact that the arrow type preserves truncation level is a direct consequence
|
The fact that the arrow type preserves truncation level is a direct consequence
|
||||||
of the fact that pi's preserve truncation level
|
of the fact that pi's preserve truncation level
|
||||||
|
|
|
@ -77,17 +77,6 @@ namespace pi
|
||||||
apply eq_of_pathover_idp, apply r
|
apply eq_of_pathover_idp, apply r
|
||||||
end
|
end
|
||||||
|
|
||||||
-- a version where C is uncurried, but where the conclusion of r is still a proper pathover
|
|
||||||
-- instead of a heterogenous equality
|
|
||||||
definition pi_pathover' {C : (Σa, B a) → Type} {f : Πb, C ⟨a, b⟩} {g : Πb', C ⟨a', b'⟩}
|
|
||||||
{p : a = a'} (r : Π(b : B a) (b' : B a') (q : b =[p] b'), f b =[dpair_eq_dpair p q] g b')
|
|
||||||
: f =[p] g :=
|
|
||||||
begin
|
|
||||||
cases p, apply pathover_idp_of_eq,
|
|
||||||
apply eq_of_homotopy, intro b,
|
|
||||||
apply (@eq_of_pathover_idp _ C), exact (r b b (pathover.idpatho b)),
|
|
||||||
end
|
|
||||||
|
|
||||||
definition pi_pathover_left {f : Πb, C a b} {g : Πb', C a' b'} {p : a = a'}
|
definition pi_pathover_left {f : Πb, C a b} {g : Πb', C a' b'} {p : a = a'}
|
||||||
(r : Π(b : B a), f b =[apo011 C p !pathover_tr] g (p ▸ b)) : f =[p] g :=
|
(r : Π(b : B a), f b =[apo011 C p !pathover_tr] g (p ▸ b)) : f =[p] g :=
|
||||||
begin
|
begin
|
||||||
|
@ -113,6 +102,36 @@ namespace pi
|
||||||
exact eq_of_pathover_idp (r b),
|
exact eq_of_pathover_idp (r b),
|
||||||
end
|
end
|
||||||
|
|
||||||
|
-- a version where C is uncurried, but where the conclusion of r is still a proper pathover
|
||||||
|
-- instead of a heterogenous equality
|
||||||
|
definition pi_pathover' {C : (Σa, B a) → Type} {f : Πb, C ⟨a, b⟩} {g : Πb', C ⟨a', b'⟩}
|
||||||
|
{p : a = a'} (r : Π(b : B a) (b' : B a') (q : b =[p] b'), f b =[dpair_eq_dpair p q] g b')
|
||||||
|
: f =[p] g :=
|
||||||
|
begin
|
||||||
|
cases p, apply pathover_idp_of_eq,
|
||||||
|
apply eq_of_homotopy, intro b,
|
||||||
|
apply (@eq_of_pathover_idp _ C), exact (r b b (pathover.idpatho b)),
|
||||||
|
end
|
||||||
|
|
||||||
|
definition pi_pathover_left' {C : (Σa, B a) → Type} {f : Πb, C ⟨a, b⟩} {g : Πb', C ⟨a', b'⟩}
|
||||||
|
{p : a = a'} (r : Π(b : B a), f b =[dpair_eq_dpair p !pathover_tr] g (p ▸ b))
|
||||||
|
: f =[p] g :=
|
||||||
|
begin
|
||||||
|
cases p, apply pathover_idp_of_eq,
|
||||||
|
apply eq_of_homotopy, intro b,
|
||||||
|
apply eq_of_pathover_idp, esimp at r, exact !pathover_ap (r b)
|
||||||
|
end
|
||||||
|
|
||||||
|
definition pi_pathover_right' {C : (Σa, B a) → Type} {f : Πb, C ⟨a, b⟩} {g : Πb', C ⟨a', b'⟩}
|
||||||
|
{p : a = a'} (r : Π(b' : B a'), f (p⁻¹ ▸ b') =[dpair_eq_dpair p !tr_pathover] g b')
|
||||||
|
: f =[p] g :=
|
||||||
|
begin
|
||||||
|
cases p, apply pathover_idp_of_eq,
|
||||||
|
apply eq_of_homotopy, intro b,
|
||||||
|
apply eq_of_pathover_idp, esimp at r, exact !pathover_ap (r b)
|
||||||
|
end
|
||||||
|
|
||||||
|
|
||||||
/- Maps on paths -/
|
/- Maps on paths -/
|
||||||
|
|
||||||
/- The action of maps given by lambda. -/
|
/- The action of maps given by lambda. -/
|
||||||
|
|
|
@ -29,10 +29,10 @@ namespace sigma
|
||||||
definition eta3 : Π (u : Σa b c, D a b c), ⟨u.1, u.2.1, u.2.2.1, u.2.2.2⟩ = u
|
definition eta3 : Π (u : Σa b c, D a b c), ⟨u.1, u.2.1, u.2.2.1, u.2.2.2⟩ = u
|
||||||
| eta3 ⟨u₁, u₂, u₃, u₄⟩ := idp
|
| eta3 ⟨u₁, u₂, u₃, u₄⟩ := idp
|
||||||
|
|
||||||
definition dpair_eq_dpair (p : a = a') (q : b =[p] b') : ⟨a, b⟩ = ⟨a', b'⟩ :=
|
definition dpair_eq_dpair [unfold 8] (p : a = a') (q : b =[p] b') : ⟨a, b⟩ = ⟨a', b'⟩ :=
|
||||||
by induction q; reflexivity
|
apo011 sigma.mk p q
|
||||||
|
|
||||||
definition sigma_eq (p : u.1 = v.1) (q : u.2 =[p] v.2) : u = v :=
|
definition sigma_eq [unfold 3 4] (p : u.1 = v.1) (q : u.2 =[p] v.2) : u = v :=
|
||||||
by induction u; induction v; exact (dpair_eq_dpair p q)
|
by induction u; induction v; exact (dpair_eq_dpair p q)
|
||||||
|
|
||||||
definition eq_pr1 [unfold 5] (p : u = v) : u.1 = v.1 :=
|
definition eq_pr1 [unfold 5] (p : u = v) : u.1 = v.1 :=
|
||||||
|
@ -157,6 +157,10 @@ namespace sigma
|
||||||
definition sigma_eq2_unc {p q : u = v} (rs : Σ(r : p..1 = q..1), p..2 =[r] q..2) : p = q :=
|
definition sigma_eq2_unc {p q : u = v} (rs : Σ(r : p..1 = q..1), p..2 =[r] q..2) : p = q :=
|
||||||
destruct rs sigma_eq2
|
destruct rs sigma_eq2
|
||||||
|
|
||||||
|
definition ap_dpair_eq_dpair (f : Πa, B a → A') (p : a = a') (q : b =[p] b')
|
||||||
|
: ap (sigma.rec f) (dpair_eq_dpair p q) = apo011 f p q :=
|
||||||
|
by induction q; reflexivity
|
||||||
|
|
||||||
/- Transport -/
|
/- Transport -/
|
||||||
|
|
||||||
/- The concrete description of transport in sigmas (and also pis) is rather trickier than in the other types. In particular, these cannot be described just in terms of transport in simpler types; they require also the dependent transport [transportD].
|
/- The concrete description of transport in sigmas (and also pis) is rather trickier than in the other types. In particular, these cannot be described just in terms of transport in simpler types; they require also the dependent transport [transportD].
|
||||||
|
|
Loading…
Reference in a new issue