feat(init.path): update init.path to use tactics, also some additions

Now the file hardly uses eq.rec explicitly anymore.
Also add the fact that horizontal and vertical inverses of paths are equal
Make one more argument explicit in eq.cancel_left and eq.cancel_right (to make it nicer to write 'apply cancel_right p')
This commit is contained in:
Floris van Doorn 2016-02-15 12:57:51 -05:00 committed by Leonardo de Moura
parent 7f09ba328e
commit ecc141779a
8 changed files with 187 additions and 155 deletions

View file

@ -51,9 +51,6 @@ namespace eq
infix ` ~2 `:50 := homotopy2
infix ` ~3 `:50 := homotopy3
definition ap011 (f : U → V → W) (Hu : u = u') (Hv : v = v') : f u v = f u' v' :=
by cases Hu; congruence; repeat assumption
definition ap0111 (f : U → V → W → X) (Hu : u = u') (Hv : v = v') (Hw : w = w')
: f u v w = f u' v' w' :=
by cases Hu; congruence; repeat assumption

View file

@ -265,7 +265,7 @@ namespace circle
{ apply equiv_pathover, intro p p' q, apply pathover_of_eq,
note H := eq_of_square (square_of_pathover q),
rewrite con_comm_base at H,
let H' := cancel_left H,
note H' := cancel_left _ H,
induction H', reflexivity}
end

View file

@ -136,7 +136,9 @@ namespace susp
refine _ ⬝ !con.assoc',
rewrite inverse2_right_inv,
refine _ ⬝ !con.assoc',
rewrite [ap_con_right_inv], unfold Susp_functor, xrewrite [idp_con_idp,-ap_compose], },
rewrite [ap_con_right_inv],
unfold Susp_functor,
xrewrite [idp_con_idp, -ap_compose (concat idp)]},
end
definition loop_susp_counit [constructor] (X : Pointed) : Susp (Ω X) →* X :=
@ -155,7 +157,7 @@ namespace susp
{ reflexivity},
{ reflexivity},
{ esimp, apply eq_pathover, apply hdeg_square,
xrewrite [ap_compose f,ap_compose (susp.elim (f x) (f x) (λ (a : f x = f x), a)),▸*],
xrewrite [ap_compose' f, ap_compose' (susp.elim (f x) (f x) (λ (a : f x = f x), a)),▸*],
xrewrite [+elim_merid,▸*,idp_con]}},
{ reflexivity}
end
@ -182,7 +184,7 @@ namespace susp
{ reflexivity},
{ exact merid pt},
{ apply eq_pathover,
xrewrite [▸*, ap_id, ap_compose (susp.elim north north (λa, a)), +elim_merid,▸*],
xrewrite [▸*, ap_id, ap_compose' (susp.elim north north (λa, a)), +elim_merid,▸*],
apply square_of_eq, exact !idp_con ⬝ !inv_con_cancel_right⁻¹}},
{ reflexivity}
end

View file

@ -150,21 +150,21 @@ namespace is_equiv
end
definition is_equiv_ap [instance] (x y : A) : is_equiv (ap f : x = y → f x = f y) :=
adjointify
(ap f)
(eq_of_fn_eq_fn' f)
(λq, !ap_con
⬝ whisker_right !ap_con _
⬝ ((!ap_inv ⬝ inverse2 (adj f _)⁻¹)
◾ (inverse (ap_compose f f⁻¹ _))
◾ (adj f _)⁻¹)
⬝ con_ap_con_eq_con_con (right_inv f) _ _
⬝ whisker_right !con.left_inv _
⬝ !idp_con)
(λp, whisker_right (whisker_left _ (ap_compose f⁻¹ f _)⁻¹) _
⬝ con_ap_con_eq_con_con (left_inv f) _ _
⬝ whisker_right !con.left_inv _
⬝ !idp_con)
adjointify
(ap f)
(eq_of_fn_eq_fn' f)
(λq, !ap_con
⬝ whisker_right !ap_con _
⬝ ((!ap_inv ⬝ inverse2 (adj f _)⁻¹)
◾ (inverse (ap_compose f f⁻¹ _))
◾ (adj f _)⁻¹)
⬝ con_ap_con_eq_con_con (right_inv f) _ _
⬝ whisker_right !con.left_inv _
⬝ !idp_con)
(λp, whisker_right (whisker_left _ (ap_compose f⁻¹ f _)⁻¹) _
⬝ con_ap_con_eq_con_con (left_inv f) _ _
⬝ whisker_right !con.left_inv _
⬝ !idp_con)
-- The function equiv_rect says that given an equivalence f : A → B,
-- and a hypothesis from B, one may always assume that the hypothesis
@ -182,17 +182,17 @@ namespace is_equiv
definition is_equiv_rect_comp (P : B → Type)
(df : Π (x : A), P (f x)) (x : A) : is_equiv_rect f P df (f x) = df x :=
calc
is_equiv_rect f P df (f x)
= right_inv f (f x) ▸ df (f⁻¹ (f x)) : by esimp
... = ap f (left_inv f x) ▸ df (f⁻¹ (f x)) : by rewrite -adj
... = left_inv f x ▸ df (f⁻¹ (f x)) : by rewrite -tr_compose
... = df x : by rewrite (apd df (left_inv f x))
calc
is_equiv_rect f P df (f x)
= right_inv f (f x) ▸ df (f⁻¹ (f x)) : by esimp
... = ap f (left_inv f x) ▸ df (f⁻¹ (f x)) : by rewrite -adj
... = left_inv f x ▸ df (f⁻¹ (f x)) : by rewrite -tr_compose
... = df x : by rewrite (apd df (left_inv f x))
theorem adj_inv (b : B) : left_inv f (f⁻¹ b) = ap f⁻¹ (right_inv f b) :=
is_equiv_rect f _
(λa,
eq.cancel_right (whisker_left _ !ap_id⁻¹ ⬝ (ap_con_eq_con_ap (left_inv f) (left_inv f a))⁻¹) ⬝
(λa, eq.cancel_right (left_inv f (id a))
(whisker_left _ !ap_id⁻¹ ⬝ (ap_con_eq_con_ap (left_inv f) (left_inv f a))⁻¹) ⬝
!ap_compose ⬝ ap02 f⁻¹ (adj f a)⁻¹)
b

View file

@ -14,7 +14,7 @@ open function eq
/- Path equality -/
namespace eq
variables {A B C : Type} {P : A → Type} {x y z t : A}
variables {A B C : Type} {P : A → Type} {a a' x y z t : A} {b b' : B}
--notation a = b := eq a b
notation x = y `:>`:50 A:49 := @eq A x y
@ -33,10 +33,10 @@ namespace eq
/- Concatenation and inverse -/
definition concat [trans] [unfold 6] (p : x = y) (q : y = z) : x = z :=
eq.rec (λp', p') q p
by induction q; exact p
definition inverse [symm] [unfold 4] (p : x = y) : y = x :=
eq.rec (refl x) p
by induction p; reflexivity
infix ⬝ := concat
postfix ⁻¹ := inverse
@ -51,57 +51,57 @@ namespace eq
-- The identity path is a right unit.
definition idp_con [unfold 4] (p : x = y) : idp ⬝ p = p :=
eq.rec_on p idp
by induction p; reflexivity
-- Concatenation is associative.
definition con.assoc' (p : x = y) (q : y = z) (r : z = t) :
p ⬝ (q ⬝ r) = (p ⬝ q) ⬝ r :=
eq.rec_on r (eq.rec_on q idp)
by induction r; induction q; reflexivity
definition con.assoc (p : x = y) (q : y = z) (r : z = t) :
(p ⬝ q) ⬝ r = p ⬝ (q ⬝ r) :=
eq.rec_on r (eq.rec_on q idp)
by induction r; induction q; reflexivity
-- The left inverse law.
definition con.right_inv [unfold 4] (p : x = y) : p ⬝ p⁻¹ = idp :=
eq.rec_on p idp
by induction p; reflexivity
-- The right inverse law.
definition con.left_inv [unfold 4] (p : x = y) : p⁻¹ ⬝ p = idp :=
eq.rec_on p idp
by induction p; reflexivity
/- Several auxiliary theorems about canceling inverses across associativity. These are somewhat
redundant, following from earlier theorems. -/
definition inv_con_cancel_left (p : x = y) (q : y = z) : p⁻¹ ⬝ (p ⬝ q) = q :=
eq.rec_on q (eq.rec_on p idp)
by induction q; induction p; reflexivity
definition con_inv_cancel_left (p : x = y) (q : x = z) : p ⬝ (p⁻¹ ⬝ q) = q :=
eq.rec_on q (eq.rec_on p idp)
by induction q; induction p; reflexivity
definition con_inv_cancel_right (p : x = y) (q : y = z) : (p ⬝ q) ⬝ q⁻¹ = p :=
eq.rec_on q (eq.rec_on p idp)
by induction q; induction p; reflexivity
definition inv_con_cancel_right (p : x = z) (q : y = z) : (p ⬝ q⁻¹) ⬝ q = p :=
eq.rec_on q (take p, eq.rec_on p idp) p
by induction q; induction p; reflexivity
-- Inverse distributes over concatenation
definition con_inv (p : x = y) (q : y = z) : (p ⬝ q)⁻¹ = q⁻¹ ⬝ p⁻¹ :=
eq.rec_on q (eq.rec_on p idp)
by induction q; induction p; reflexivity
definition inv_con_inv_left (p : y = x) (q : y = z) : (p⁻¹ ⬝ q)⁻¹ = q⁻¹ ⬝ p :=
eq.rec_on q (eq.rec_on p idp)
by induction q; induction p; reflexivity
-- universe metavariables
definition inv_con_inv_right (p : x = y) (q : z = y) : (p ⬝ q⁻¹)⁻¹ = q ⬝ p⁻¹ :=
eq.rec_on p (take q, eq.rec_on q idp) q
by induction q; induction p; reflexivity
definition inv_con_inv_inv (p : y = x) (q : z = y) : (p⁻¹ ⬝ q⁻¹)⁻¹ = q ⬝ p :=
eq.rec_on p (eq.rec_on q idp)
by induction q; induction p; reflexivity
-- Inverse is an involution.
definition inv_inv (p : x = y) : p⁻¹⁻¹ = p :=
eq.rec_on p idp
by induction p; reflexivity
-- auxiliary definition used by 'cases' tactic
definition elim_inv_inv {A : Type} {a b : A} {C : a = b → Type} (H₁ : a = b) (H₂ : C (H₁⁻¹⁻¹)) : C H₁ :=
@ -111,59 +111,61 @@ namespace eq
definition con_eq_of_eq_inv_con {p : x = z} {q : y = z} {r : y = x} :
p = r⁻¹ ⬝ q → r ⬝ p = q :=
eq.rec_on r (take p h, !idp_con ⬝ h ⬝ !idp_con) p
begin
induction r, intro h, exact !idp_con ⬝ h ⬝ !idp_con
end
definition con_eq_of_eq_con_inv [unfold 5] {p : x = z} {q : y = z} {r : y = x} :
r = q ⬝ p⁻¹ → r ⬝ p = q :=
eq.rec_on p (take q h, h) q
by induction p; exact id
definition inv_con_eq_of_eq_con {p : x = z} {q : y = z} {r : x = y} :
p = r ⬝ q → r⁻¹ ⬝ p = q :=
eq.rec_on r (take q h, !idp_con ⬝ h ⬝ !idp_con) q
by induction r; intro h; exact !idp_con ⬝ h ⬝ !idp_con
definition con_inv_eq_of_eq_con [unfold 5] {p : z = x} {q : y = z} {r : y = x} :
r = q ⬝ p → r ⬝ p⁻¹ = q :=
eq.rec_on p (take r h, h) r
by induction p; exact id
definition eq_con_of_inv_con_eq {p : x = z} {q : y = z} {r : y = x} :
r⁻¹ ⬝ q = p → q = r ⬝ p :=
eq.rec_on r (take p h, !idp_con⁻¹ ⬝ h ⬝ !idp_con⁻¹) p
by induction r; intro h; exact !idp_con⁻¹ ⬝ h ⬝ !idp_con⁻¹
definition eq_con_of_con_inv_eq [unfold 5] {p : x = z} {q : y = z} {r : y = x} :
q ⬝ p⁻¹ = r → q = r ⬝ p :=
eq.rec_on p (take q h, h) q
by induction p; exact id
definition eq_inv_con_of_con_eq {p : x = z} {q : y = z} {r : x = y} :
r ⬝ q = p → q = r⁻¹ ⬝ p :=
eq.rec_on r (take q h, !idp_con⁻¹ ⬝ h ⬝ !idp_con⁻¹) q
by induction r; intro h; exact !idp_con⁻¹ ⬝ h ⬝ !idp_con⁻¹
definition eq_con_inv_of_con_eq [unfold 5] {p : z = x} {q : y = z} {r : y = x} :
q ⬝ p = r → q = r ⬝ p⁻¹ :=
eq.rec_on p (take r h, h) r
by induction p; exact id
definition eq_of_con_inv_eq_idp [unfold 5] {p q : x = y} : p ⬝ q⁻¹ = idp → p = q :=
eq.rec_on q (take p h, h) p
by induction q; exact id
definition eq_of_inv_con_eq_idp {p q : x = y} : q⁻¹ ⬝ p = idp → p = q :=
eq.rec_on q (take p h, !idp_con⁻¹ ⬝ h) p
by induction q; intro h; exact !idp_con⁻¹ ⬝ h
definition eq_inv_of_con_eq_idp' [unfold 5] {p : x = y} {q : y = x} : p ⬝ q = idp → p = q⁻¹ :=
eq.rec_on q (take p h, h) p
by induction q; exact id
definition eq_inv_of_con_eq_idp {p : x = y} {q : y = x} : q ⬝ p = idp → p = q⁻¹ :=
eq.rec_on q (take p h, !idp_con⁻¹ ⬝ h) p
by induction q; intro h; exact !idp_con⁻¹ ⬝ h
definition eq_of_idp_eq_inv_con {p q : x = y} : idp = p⁻¹ ⬝ q → p = q :=
eq.rec_on p (take q h, h ⬝ !idp_con) q
by induction p; intro h; exact h ⬝ !idp_con
definition eq_of_idp_eq_con_inv [unfold 4] {p q : x = y} : idp = q ⬝ p⁻¹ → p = q :=
eq.rec_on p (take q h, h) q
by induction p; exact id
definition inv_eq_of_idp_eq_con [unfold 4] {p : x = y} {q : y = x} : idp = q ⬝ p → p⁻¹ = q :=
eq.rec_on p (take q h, h) q
by induction p; exact id
definition inv_eq_of_idp_eq_con' {p : x = y} {q : y = x} : idp = p ⬝ q → p⁻¹ = q :=
eq.rec_on p (take q h, h ⬝ !idp_con) q
by induction p; intro h; exact h ⬝ !idp_con
definition con_inv_eq_idp [unfold 6] {p q : x = y} (r : p = q) : p ⬝ q⁻¹ = idp :=
by cases r;apply con.right_inv
@ -187,7 +189,7 @@ namespace eq
definition transport [subst] [reducible] [unfold 5] (P : A → Type) {x y : A} (p : x = y)
(u : P x) : P y :=
eq.rec_on p u
by induction p; exact u
-- This idiom makes the operation right associative.
infixr ` ▸ ` := transport _
@ -203,7 +205,7 @@ namespace eq
p⁻¹ ▸ u
definition ap [unfold 6] ⦃A B : Type⦄ (f : A → B) {x y:A} (p : x = y) : f x = f y :=
eq.rec_on p idp
by induction p; reflexivity
abbreviation ap01 [parsing_only] := ap
@ -227,42 +229,45 @@ namespace eq
H1 ▸ homotopy.refl f
definition apd10 [unfold 5] {f g : Πx, P x} (H : f = g) : f ~ g :=
λx, eq.rec_on H idp
λx, by induction H; reflexivity
--the next theorem is useful if you want to write "apply (apd10' a)"
definition apd10' [unfold 6] {f g : Πx, P x} (a : A) (H : f = g) : f a = g a :=
eq.rec_on H idp
by induction H; reflexivity
--apd10 is also ap evaluation
definition apd10_eq_ap_eval {f g : Πx, P x} (H : f = g) (a : A)
: apd10 H a = ap (λs : Πx, P x, s a) H :=
eq.rec_on H idp
by induction H; reflexivity
definition ap10 [reducible] [unfold 5] {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 :=
eq.rec_on H (eq.rec_on p idp)
by induction H; exact ap f p
definition apd [unfold 6] (f : Πa, P a) {x y : A} (p : x = y) : p ▸ f x = f y :=
eq.rec_on p idp
by induction p; reflexivity
definition ap011 (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' :=
by cases Ha; exact ap (f a) Hb
/- More theorems for moving things around in equations -/
definition tr_eq_of_eq_inv_tr {P : A → Type} {x y : A} {p : x = y} {u : P x} {v : P y} :
u = p⁻¹ ▸ v → p ▸ u = v :=
eq.rec_on p (take v, id) v
by induction p; exact id
definition inv_tr_eq_of_eq_tr {P : A → Type} {x y : A} {p : y = x} {u : P x} {v : P y} :
u = p ▸ v → p⁻¹ ▸ u = v :=
eq.rec_on p (take u, id) u
by induction p; exact id
definition eq_inv_tr_of_tr_eq {P : A → Type} {x y : A} {p : x = y} {u : P x} {v : P y} :
p ▸ u = v → u = p⁻¹ ▸ v :=
eq.rec_on p (take v, id) v
by induction p; exact id
definition eq_tr_of_inv_tr_eq {P : A → Type} {x y : A} {p : y = x} {u : P x} {v : P y} :
p⁻¹ ▸ u = v → u = p ▸ v :=
eq.rec_on p (take u, id) u
by induction p; exact id
/- Functoriality of functions -/
@ -275,104 +280,102 @@ namespace eq
-- Functions commute with concatenation.
definition ap_con [unfold 8] (f : A → B) {x y z : A} (p : x = y) (q : y = z) :
ap f (p ⬝ q) = ap f p ⬝ ap f q :=
eq.rec_on q idp
by induction q; reflexivity
definition con_ap_con_eq_con_ap_con_ap (f : A → B) {w x y z : A} (r : f w = f x)
(p : x = y) (q : y = z) : r ⬝ ap f (p ⬝ q) = (r ⬝ ap f p) ⬝ ap f q :=
eq.rec_on q (take p, eq.rec_on p idp) p
by induction q; induction p; reflexivity
definition ap_con_con_eq_ap_con_ap_con (f : A → B) {w x y z : A} (p : x = y) (q : y = z)
(r : f z = f w) : ap f (p ⬝ q) ⬝ r = ap f p ⬝ (ap f q ⬝ r) :=
eq.rec_on q (eq.rec_on p (take r, con.assoc _ _ _)) r
by induction q; induction p; apply con.assoc
-- Functions commute with path inverses.
definition ap_inv' [unfold 6] (f : A → B) {x y : A} (p : x = y) : (ap f p)⁻¹ = ap f p⁻¹ :=
eq.rec_on p idp
by induction p; reflexivity
definition ap_inv [unfold 6] (f : A → B) {x y : A} (p : x = y) : ap f p⁻¹ = (ap f p)⁻¹ :=
eq.rec_on p idp
by induction p; reflexivity
-- [ap] itself is functorial in the first argument.
definition ap_id [unfold 4] (p : x = y) : ap id p = p :=
eq.rec_on p idp
by induction p; reflexivity
definition ap_compose [unfold 8] (g : B → C) (f : A → B) {x y : A} (p : x = y) :
ap (g ∘ f) p = ap g (ap f p) :=
eq.rec_on p idp
by induction p; reflexivity
-- Sometimes we don't have the actual function [compose].
definition ap_compose' [unfold 8] (g : B → C) (f : A → B) {x y : A} (p : x = y) :
ap (λa, g (f a)) p = ap g (ap f p) :=
eq.rec_on p idp
by induction p; reflexivity
-- The action of constant maps.
definition ap_constant [unfold 5] (p : x = y) (z : B) : ap (λu, z) p = idp :=
eq.rec_on p idp
by induction p; reflexivity
-- Naturality of [ap].
-- see also natural_square in cubical.square
definition ap_con_eq_con_ap {f g : A → B} (p : f ~ g) {x y : A} (q : x = y) :
ap f q ⬝ p y = p x ⬝ ap g q :=
eq.rec_on q !idp_con
by induction q; apply idp_con
-- Naturality of [ap] at identity.
definition ap_con_eq_con {f : A → A} (p : Πx, f x = x) {x y : A} (q : x = y) :
ap f q ⬝ p y = p x ⬝ q :=
eq.rec_on q !idp_con
by induction q; apply idp_con
definition con_ap_eq_con {f : A → A} (p : Πx, x = f x) {x y : A} (q : x = y) :
p x ⬝ ap f q = q ⬝ p y :=
eq.rec_on q !idp_con⁻¹
by induction q; exact !idp_con⁻¹
-- Naturality of [ap] with constant function
definition ap_con_eq {f : A → B} {b : B} (p : Πx, f x = b) {x y : A} (q : x = y) :
ap f q ⬝ p y = p x :=
eq.rec_on q !idp_con
by induction q; apply idp_con
-- Naturality with other paths hanging around.
definition con_ap_con_con_eq_con_con_ap_con {f g : A → B} (p : f ~ g) {x y : A} (q : x = y)
{w z : B} (r : w = f x) (s : g y = z) :
(r ⬝ ap f q) ⬝ (p y ⬝ s) = (r ⬝ p x) ⬝ (ap g q ⬝ s) :=
eq.rec_on s (eq.rec_on q idp)
by induction s; induction q; reflexivity
definition con_ap_con_eq_con_con_ap {f g : A → B} (p : f ~ g) {x y : A} (q : x = y)
{w : B} (r : w = f x) :
(r ⬝ ap f q) ⬝ p y = (r ⬝ p x) ⬝ ap g q :=
eq.rec_on q idp
by induction q; reflexivity
-- TODO: try this using the simplifier, and compare proofs
definition ap_con_con_eq_con_ap_con {f g : A → B} (p : f ~ g) {x y : A} (q : x = y)
{z : B} (s : g y = z) :
ap f q ⬝ (p y ⬝ s) = p x ⬝ (ap g q ⬝ s) :=
eq.rec_on s (eq.rec_on q
(calc
(ap f idp) ⬝ (p x ⬝ idp) = idp ⬝ p x : idp
... = p x : !idp_con
... = (p x) ⬝ (ap g idp ⬝ idp) : idp))
-- This also works:
-- eq.rec_on s (eq.rec_on q (!idp_con ▸ idp))
begin
induction s,
induction q,
apply idp_con
end
definition con_ap_con_con_eq_con_con_con {f : A → A} (p : f ~ id) {x y : A} (q : x = y)
{w z : A} (r : w = f x) (s : y = z) :
(r ⬝ ap f q) ⬝ (p y ⬝ s) = (r ⬝ p x) ⬝ (q ⬝ s) :=
eq.rec_on s (eq.rec_on q idp)
by induction s; induction q; reflexivity
definition con_con_ap_con_eq_con_con_con {g : A → A} (p : id ~ g) {x y : A} (q : x = y)
{w z : A} (r : w = x) (s : g y = z) :
(r ⬝ p x) ⬝ (ap g q ⬝ s) = (r ⬝ q) ⬝ (p y ⬝ s) :=
eq.rec_on s (eq.rec_on q idp)
by induction s; induction q; reflexivity
definition con_ap_con_eq_con_con {f : A → A} (p : f ~ id) {x y : A} (q : x = y)
{w : A} (r : w = f x) :
(r ⬝ ap f q) ⬝ p y = (r ⬝ p x) ⬝ q :=
eq.rec_on q idp
by induction q; reflexivity
definition ap_con_con_eq_con_con {f : A → A} (p : f ~ id) {x y : A} (q : x = y)
{z : A} (s : y = z) :
ap f q ⬝ (p y ⬝ s) = p x ⬝ (q ⬝ s) :=
eq.rec_on s (eq.rec_on q (!idp_con ▸ idp))
by induction s; induction q; apply idp_con
definition con_con_ap_eq_con_con {g : A → A} (p : id ~ g) {x y : A} (q : x = y)
{w : A} (r : w = x) :
@ -382,11 +385,7 @@ namespace eq
definition con_ap_con_eq_con_con' {g : A → A} (p : id ~ g) {x y : A} (q : x = y)
{z : A} (s : g y = z) :
p x ⬝ (ap g q ⬝ s) = q ⬝ (p y ⬝ s) :=
begin
apply (eq.rec_on s),
apply (eq.rec_on q),
apply (idp_con (p x) ▸ idp)
end
by induction s; induction q; exact !idp_con⁻¹
/- Action of [apd10] and [ap10] on paths -/
@ -396,11 +395,11 @@ namespace eq
definition apd10_con {f f' f'' : Πx, P x} (h : f = f') (h' : f' = f'') (x : A) :
apd10 (h ⬝ h') x = apd10 h x ⬝ apd10 h' x :=
eq.rec_on h (take h', eq.rec_on h' idp) h'
by induction h; induction h'; reflexivity
definition apd10_inv {f g : Πx : A, P x} (h : f = g) (x : A) :
apd10 h⁻¹ x = (apd10 h x)⁻¹ :=
eq.rec_on h idp
by induction h; reflexivity
definition ap10_idp {f : A → B} (x : A) : ap10 (refl f) x = idp := idp
@ -413,7 +412,7 @@ namespace eq
-- [ap10] also behaves nicely on paths produced by [ap]
definition ap_ap10 (f g : A → B) (h : B → C) (p : f = g) (a : A) :
ap h (ap10 p a) = ap10 (ap (λ f', h ∘ f') p) a:=
eq.rec_on p idp
by induction p; reflexivity
/- Transport and the groupoid structure of paths -/
@ -422,7 +421,7 @@ namespace eq
definition con_tr [unfold 7] {P : A → Type} {x y z : A} (p : x = y) (q : y = z) (u : P x) :
p ⬝ q ▸ u = q ▸ p ▸ u :=
eq.rec_on q idp
by induction q; reflexivity
definition tr_inv_tr {P : A → Type} {x y : A} (p : x = y) (z : P y) :
p ▸ p⁻¹ ▸ z = z :=
@ -438,28 +437,30 @@ namespace eq
ap (transport P r) (con_tr p q u)
= (con_tr p (q ⬝ r) u) ⬝ (con_tr q r (p ▸ u))
:> ((p ⬝ (q ⬝ r)) ▸ u = r ▸ q ▸ p ▸ u) :=
eq.rec_on r (eq.rec_on q (eq.rec_on p idp))
by induction r; induction q; induction p; reflexivity
-- Here is another coherence lemma for transport.
definition tr_inv_tr_lemma {P : A → Type} {x y : A} (p : x = y) (z : P x) :
tr_inv_tr p (transport P p z) = ap (transport P p) (inv_tr_tr p z) :=
eq.rec_on p idp
by induction p; reflexivity
/- some properties for apd -/
definition apd_idp (x : A) (f : Πx, P x) : apd f idp = idp :> (f x = f x) := idp
definition apd_con (f : Πx, P x) {x y z : A} (p : x = y) (q : y = z)
: apd f (p ⬝ q) = con_tr p q (f x) ⬝ ap (transport P q) (apd f p) ⬝ apd f q :=
by cases p;cases q;apply idp
by cases p; cases q; apply idp
definition apd_inv (f : Πx, P x) {x y : A} (p : x = y)
: apd f p⁻¹ = (eq_inv_tr_of_tr_eq (apd f p))⁻¹ :=
by cases p;apply idp
by cases p; apply idp
-- Dependent transport in a doubly dependent type.
definition transportD [unfold 6] {P : A → Type} (Q : Πa, P a → Type)
{a a' : A} (p : a = a') (b : P a) (z : Q a b) : Q a' (p ▸ b) :=
eq.rec_on p z
by induction p; exact z
-- In Coq the variables P, Q and b are explicit, but in Lean we can probably have them implicit
-- using the following notation
@ -473,34 +474,33 @@ namespace eq
notation p ` ▸2 `:65 x:64 := transport2 _ p _ x
-- An alternative definition.
definition tr2_eq_ap10 (Q : A → Type) {x y : A} {p q : x = y} (r : p = q)
(z : Q x) :
definition tr2_eq_ap10 (Q : A → Type) {x y : A} {p q : x = y} (r : p = q) (z : Q x) :
transport2 Q r z = ap10 (ap (transport Q) r) z :=
eq.rec_on r idp
by induction r; reflexivity
definition tr2_con {P : A → Type} {x y : A} {p1 p2 p3 : x = y}
(r1 : p1 = p2) (r2 : p2 = p3) (z : P x) :
transport2 P (r1 ⬝ r2) z = transport2 P r1 z ⬝ transport2 P r2 z :=
eq.rec_on r1 (eq.rec_on r2 idp)
by induction r1; induction r2; reflexivity
definition tr2_inv (Q : A → Type) {x y : A} {p q : x = y} (r : p = q) (z : Q x) :
transport2 Q r⁻¹ z = (transport2 Q r z)⁻¹ :=
eq.rec_on r idp
by induction r; reflexivity
definition transportD2 [unfold 7] (B C : A → Type) (D : Π(a:A), B a → C a → Type)
definition transportD2 [unfold 7] {B C : A → Type} (D : Π(a:A), B a → C a → Type)
{x1 x2 : A} (p : x1 = x2) (y : B x1) (z : C x1) (w : D x1 y z) : D x2 (p ▸ y) (p ▸ z) :=
eq.rec_on p w
by induction p; exact w
notation p ` ▸D2 `:65 x:64 := transportD2 _ _ _ p _ _ x
notation p ` ▸D2 `:65 x:64 := transportD2 _ p _ _ x
definition ap_tr_con_tr2 (P : A → Type) {x y : A} {p q : x = y} {z w : P x} (r : p = q)
(s : z = w) :
ap (transport P p) s ⬝ transport2 P r w = transport2 P r z ⬝ ap (transport P q) s :=
eq.rec_on r !idp_con⁻¹
by induction r; exact !idp_con⁻¹
definition fn_tr_eq_tr_fn {P Q : A → Type} {x y : A} (p : x = y) (f : Πx, P x → Q x) (z : P x) :
f y (p ▸ z) = (p ▸ (f x z)) :=
eq.rec_on p idp
by induction p; reflexivity
/- Transporting in particular fibrations -/
@ -515,47 +515,47 @@ namespace eq
-- Transporting in a constant fibration.
definition tr_constant (p : x = y) (z : B) : transport (λx, B) p z = z :=
eq.rec_on p idp
by induction p; reflexivity
definition tr2_constant {p q : x = y} (r : p = q) (z : B) :
tr_constant p z = transport2 (λu, B) r z ⬝ tr_constant q z :=
eq.rec_on r !idp_con⁻¹
by induction r; exact !idp_con⁻¹
-- Transporting in a pulled back fibration.
definition tr_compose (P : B → Type) (f : A → B) (p : x = y) (z : P (f x)) :
transport (P ∘ f) p z = transport P (ap f p) z :=
eq.rec_on p idp
by induction p; reflexivity
definition ap_precompose (f : A → B) (g g' : B → C) (p : g = g') :
ap (λh, h ∘ f) p = transport (λh : B → C, g ∘ f = h ∘ f) p idp :=
eq.rec_on p idp
by induction p; reflexivity
definition apd10_ap_precompose (f : A → B) (g g' : B → C) (p : g = g') :
apd10 (ap (λh : B → C, h ∘ f) p) = λa, apd10 p (f a) :=
eq.rec_on p idp
by induction p; reflexivity
definition apd10_ap_precompose_dependent {C : B → Type}
(f : A → B) {g g' : Πb : B, C b} (p : g = g')
: apd10 (ap (λ(h : (Πb : B, C b))(a : A), h (f a)) p) = λa, apd10 p (f a) :=
eq.rec_on p idp
by induction p; reflexivity
definition apd10_ap_postcompose (f : B → C) (g g' : A → B) (p : g = g') :
apd10 (ap (λh : A → B, f ∘ h) p) = λa, ap f (apd10 p a) :=
eq.rec_on p idp
by induction p; reflexivity
-- A special case of [tr_compose] which seems to come up a lot.
definition tr_eq_cast_ap {P : A → Type} {x y} (p : x = y) (u : P x) : p ▸ u = cast (ap P p) u :=
eq.rec_on p idp
by induction p; reflexivity
definition tr_eq_cast_ap_fn {P : A → Type} {x y} (p : x = y) : transport P p = cast (ap P p) :=
eq.rec_on p idp
by induction p; reflexivity
/- The behavior of [ap] and [apd] -/
-- In a constant fibration, [apd] reduces to [ap], modulo [transport_const].
definition apd_eq_tr_constant_con_ap (f : A → B) (p : x = y) :
apd f p = tr_constant p (f x) ⬝ ap f p :=
eq.rec_on p idp
by induction p; reflexivity
/- The 2-dimensional groupoid structure -/
@ -563,7 +563,7 @@ namespace eq
-- Horizontal composition of 2-dimensional paths.
definition concat2 [unfold 9 10] {p p' : x = y} {q q' : y = z} (h : p = p') (h' : q = q')
: p ⬝ q = p' ⬝ q' :=
eq.rec_on h (eq.rec_on h' idp)
ap011 concat h h'
-- 2-dimensional path inversion
definition inverse2 [unfold 6] {p q : x = y} (h : p = q) : p⁻¹ = q⁻¹ :=
@ -582,17 +582,17 @@ namespace eq
-- Unwhiskering, a.k.a. cancelling
definition cancel_left {x y z : A} {p : x = y} {q r : y = z} : (p ⬝ q = p ⬝ r) → (q = r) :=
definition cancel_left {x y z : A} (p : x = y) {q r : y = z} : (p ⬝ q = p ⬝ r) → (q = r) :=
λs, !inv_con_cancel_left⁻¹ ⬝ whisker_left p⁻¹ s ⬝ !inv_con_cancel_left
definition cancel_right {x y z : A} {p q : x = y} {r : y = z} : (p ⬝ r = q ⬝ r) → (p = q) :=
definition cancel_right {x y z : A} {p q : x = y} (r : y = z) : (p ⬝ r = q ⬝ r) → (p = q) :=
λs, !con_inv_cancel_right⁻¹ ⬝ whisker_right s r⁻¹ ⬝ !con_inv_cancel_right
-- Whiskering and identity paths.
definition whisker_right_idp {p q : x = y} (h : p = q) :
whisker_right h idp = h :=
eq.rec_on h (eq.rec_on p idp)
by induction h; induction p; reflexivity
definition whisker_right_idp_left [unfold_full] (p : x = y) (q : y = z) :
whisker_right idp q = idp :> (p ⬝ q = p ⬝ q) :=
@ -603,8 +603,15 @@ namespace eq
idp
definition whisker_left_idp {p q : x = y} (h : p = q) :
(idp_con p) ⁻¹ ⬝ whisker_left idp h ⬝ idp_con q = h :=
eq.rec_on h (eq.rec_on p idp)
(idp_con p)⁻¹ ⬝ whisker_left idp h ⬝ idp_con q = h :=
by induction h; induction p; reflexivity
definition whisker_left_idp2 {A : Type} {a : A} (p : idp = idp :> a = a) :
whisker_left idp p = p :=
begin
refine _ ⬝ whisker_left_idp p,
exact !idp_con⁻¹
end
definition con2_idp [unfold_full] {p q : x = y} (h : p = q) :
h ◾ idp = whisker_right h idp :> (p ⬝ idp = q ⬝ idp) :=
@ -614,16 +621,28 @@ namespace eq
idp ◾ h = whisker_left idp h :> (idp ⬝ p = idp ⬝ q) :=
idp
definition inverse2_concat2 {p p' : x = y} (h : p = p')
: h⁻² ◾ h = con.left_inv p ⬝ (con.left_inv p')⁻¹ :=
by induction h; induction p; reflexivity
-- The interchange law for concatenation.
definition con2_con_con2 {p p' p'' : x = y} {q q' q'' : y = z}
(a : p = p') (b : p' = p'') (c : q = q') (d : q' = q'') :
(a ◾ c) ⬝ (b ◾ d) = (a ⬝ b) ◾ (c ⬝ d) :=
eq.rec_on d (eq.rec_on c (eq.rec_on b (eq.rec_on a idp)))
by induction d; induction c; induction b;induction a; reflexivity
definition concat2_eq_rl {A : Type} {x y z : A} {p p' : x = y} {q q' : y = z}
(a : p = p') (b : q = q') : a ◾ b = whisker_right a q ⬝ whisker_left p' b :=
by induction b; induction a; reflexivity
definition concat2_eq_lf {A : Type} {x y z : A} {p p' : x = y} {q q' : y = z}
(a : p = p') (b : q = q') : a ◾ b = whisker_left p b ⬝ whisker_right a q' :=
by induction b; induction a; reflexivity
definition whisker_right_con_whisker_left {x y z : A} {p p' : x = y} {q q' : y = z}
(a : p = p') (b : q = q') :
(whisker_right a q) ⬝ (whisker_left p' b) = (whisker_left p b) ⬝ (whisker_right a q') :=
eq.rec_on b (eq.rec_on a !idp_con⁻¹)
by induction b; induction a; reflexivity
-- Structure corresponding to the coherence equations of a bicategory.
@ -638,14 +657,28 @@ namespace eq
-- The 3-cell witnessing the left unit triangle.
definition triangulator (p : x = y) (q : y = z) :
con.assoc' p idp q ⬝ whisker_right (con_idp p) q = whisker_left p (idp_con q) :=
eq.rec_on q (eq.rec_on p idp)
by induction q; induction p; reflexivity
definition eckmann_hilton {x:A} (p q : idp = idp :> x = x) : p ⬝ q = q ⬝ p :=
(!whisker_right_idp ◾ !whisker_left_idp)⁻¹
⬝ whisker_left _ !idp_con
⬝ !whisker_right_con_whisker_left
⬝ whisker_right !idp_con⁻¹ _
⬝ (!whisker_left_idp ◾ !whisker_right_idp)
begin
refine (whisker_right_idp p ◾ whisker_left_idp2 q)⁻¹ ⬝ _,
refine !whisker_right_con_whisker_left ⬝ _,
refine !whisker_left_idp2 ◾ !whisker_right_idp
end
definition concat_eq_concat2 {A : Type} {a : A} (p q : idp = idp :> a = a) : p ⬝ q = p ◾ q :=
begin
refine (whisker_right_idp p ◾ whisker_left_idp2 q)⁻¹ ⬝ _,
exact !concat2_eq_rl⁻¹
end
definition inverse_eq_inverse2 {A : Type} {a : A} (p : idp = idp :> a = a) : p⁻¹ = p⁻² :=
begin
apply eq.cancel_right p,
refine !con.left_inv ⬝ _,
refine _ ⬝ !concat_eq_concat2⁻¹,
exact !inverse2_concat2⁻¹,
end
-- The action of functions on 2-dimensional paths
definition ap02 [unfold 8] [reducible] (f : A → B) {x y : A} {p q : x = y} (r : p = q)
@ -654,18 +687,18 @@ namespace eq
definition ap02_con (f : A → B) {x y : A} {p p' p'' : x = y} (r : p = p') (r' : p' = p'') :
ap02 f (r ⬝ r') = ap02 f r ⬝ ap02 f r' :=
eq.rec_on r (eq.rec_on r' idp)
by induction r; induction r'; reflexivity
definition ap02_con2 (f : A → B) {x y z : A} {p p' : x = y} {q q' :y = z} (r : p = p')
(s : q = q') :
ap02 f (r ◾ s) = ap_con f p q
⬝ (ap02 f r ◾ ap02 f s)
⬝ (ap_con f p' q')⁻¹ :=
eq.rec_on r (eq.rec_on s (eq.rec_on q (eq.rec_on p idp)))
by induction r; induction s; induction q; induction p; reflexivity
definition apd02 [unfold 8] {p q : x = y} (f : Π x, P x) (r : p = q) :
apd f p = transport2 P r (f x) ⬝ apd f q :=
eq.rec_on r !idp_con⁻¹
by induction r; exact !idp_con⁻¹
-- And now for a lemma whose statement is much longer than its proof.
definition apd02_con {P : A → Type} (f : Π x:A, P x) {x y : A}
@ -674,6 +707,6 @@ namespace eq
⬝ whisker_left (transport2 P r1 (f x)) (apd02 f r2)
⬝ con.assoc' _ _ _
⬝ (whisker_right (tr2_con r1 r2 (f x))⁻¹ (apd f p3)) :=
eq.rec_on r2 (eq.rec_on r1 (eq.rec_on p1 idp))
by induction r2; induction r1; induction p1; reflexivity
end eq

View file

@ -171,7 +171,7 @@ namespace pointed
begin
fconstructor,
{ intro a, reflexivity},
{ esimp, exact !idp_con ⬝ !ap_id⁻¹}
{ reflexivity}
end
definition comp_pid (f : A →* B) : f ∘* pid A ~* f :=

View file

@ -104,7 +104,7 @@ namespace is_trunc
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : transport_eq_r
... = imp (transport (λx, R a x) p (refl a)) : H3
... = imp (refl a) : is_hprop.elim,
cancel_left H4)
cancel_left (imp (refl a)) H4)
definition relation_equiv_eq {A : Type} (R : A → A → Type)
(mere : Π(a b : A), is_hprop (R a b)) (refl : Π(a : A), R a a)

View file

@ -37,7 +37,7 @@
"" "" "" "" "×" "Σ" "Π" "~" "||" "&&" "" "" ""
"" "" "" "" "" "𝔸"
;; HoTT notation
"Ω" "" "map₊" "" "π₁" "" "" "" "" ""
"Ω" "" "map₊" "" "π₁" "" "S¹." "" "" "" ""
"⁻¹ᵉ" "⁻¹ᶠ" "⁻¹ᵍ" "⁻¹ʰ" "⁻¹ⁱ" "⁻¹ᵐ" "⁻¹ᵒ" "⁻¹ᵖ" "⁻¹ʳ" "⁻¹ᵛ" "⁻¹ˢ" "⁻²" "⁻²ᵒ"
"⬝e" "⬝i" "⬝o" "⬝op" "⬝po" "⬝h" "⬝v" "⬝hp" "⬝vp" "⬝ph" "⬝pv" "⬝r" "" "◾o"
"∘n" "∘f" "∘fi" "∘nf" "∘fn" "∘n1f" "∘1nf" "∘f1n" "∘fn1"