style(hott): let inverse notation have higher binding power than application

This commit is contained in:
Floris van Doorn 2015-02-27 19:02:18 -05:00
parent 219f7ae11a
commit 23a248ab28
19 changed files with 227 additions and 227 deletions

View file

@ -42,6 +42,9 @@ prefix `-` := has_neg.neg
notation 1 := !has_one.one
notation 0 := !has_zero.zero
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ᵍ`:std.prec.max_plus := has_inv.inv
/- semigroup -/
@ -425,7 +428,7 @@ section add_group
theorem zero_sub (a : A) : 0 - a = -a := !zero_add
theorem sub_zero (a : A) : a - 0 = a := (neg_zero⁻¹) ▹ !add_zero
theorem sub_zero (a : A) : a - 0 = a := neg_zero⁻¹ ▹ !add_zero
theorem sub_neg_eq_add (a b : A) : a - (-b) = a + b := !neg_neg ▹ idp

View file

@ -34,7 +34,7 @@ namespace category
(λ (a b c d : A) (p : c = d) (q : b = c) (r : a = b), con.assoc r q p)
(λ (a b : A) (p : a = b), con_idp p)
(λ (a b : A) (p : a = b), idp_con p)
(λ (a b : A) (p : a = b), @is_iso.mk A _ a b p (p⁻¹)
(λ (a b : A) (p : a = b), @is_iso.mk A _ a b p p⁻¹
!con.left_inv !con.right_inv)
-- A groupoid with a contractible carrier is a group
@ -44,7 +44,7 @@ namespace category
fapply group.mk,
intros (f, g), apply (comp f g),
apply homH,
intros (f, g, h), apply ((assoc f g h)⁻¹),
intros (f, g, h), apply (assoc f g h)⁻¹,
apply (ID (center ob)),
intro f, apply id_left,
intro f, apply id_right,
@ -57,7 +57,7 @@ namespace category
fapply group.mk,
intros (f, g), apply (comp f g),
apply homH,
intros (f, g, h), apply ((assoc f g h)⁻¹),
intros (f, g, h), apply (assoc f g h)⁻¹,
apply (ID ⋆),
intro f, apply id_left,
intro f, apply id_right,
@ -73,7 +73,7 @@ namespace category
intros, apply (@group.carrier_hset A G),
intros (a, b, c, g, h), exact (@group.mul A G g h),
intro a, exact (@group.one A G),
intros, exact ((@group.mul_assoc A G h g f)⁻¹),
intros, exact (@group.mul_assoc A G h g f)⁻¹,
intros, exact (@group.one_mul A G f),
intros, exact (@group.mul_one A G f),
intros, apply is_iso.mk,
@ -87,7 +87,7 @@ namespace category
fapply group.mk,
intros (f, g), apply (comp f g),
apply homH,
intros (f, g, h), apply ((assoc f g h)⁻¹),
intros (f, g, h), apply (assoc f g h)⁻¹,
apply (ID a),
intro f, apply id_left,
intro f, apply id_right,

View file

@ -2,148 +2,139 @@
Copyright (c) 2014 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Module: algebra.precategory.morphism
Authors: Floris van Doorn, Jakob von Raumer
Module: algebra.category.morphism
Author: Floris van Doorn, Jakob von Raumer
-/
import .basic
import algebra.precategory.basic
open eq category sigma sigma.ops equiv is_equiv function is_trunc
open eq category sigma sigma.ops equiv is_equiv is_trunc
namespace morphism
structure is_section [class] {ob : Type} [C : precategory ob] {a b : ob} (f : a ⟶ b) :=
{retraction_of : b ⟶ a}
(retraction_compose : retraction_of ∘ f = id)
structure is_retraction [class] {ob : Type} [C : precategory ob] {a b : ob} (f : a ⟶ b) :=
{section_of : b ⟶ a}
(compose_section : f ∘ section_of = id)
structure is_iso [class] {ob : Type} [C : precategory ob] {a b : ob} (f : a ⟶ b) :=
{inverse : b ⟶ a}
(inverse_compose : inverse ∘ f = id)
(compose_inverse : f ∘ inverse = id)
attribute is_iso [multiple-instances]
open is_section is_retraction is_iso
definition retraction_of [reducible] := @is_section.retraction_of
definition retraction_compose [reducible] := @is_section.retraction_compose
definition section_of [reducible] := @is_retraction.section_of
definition compose_section [reducible] := @is_retraction.compose_section
definition inverse [reducible] := @is_iso.inverse
definition inverse_compose [reducible] := @is_iso.inverse_compose
definition compose_inverse [reducible] := @is_iso.compose_inverse
postfix `⁻¹` := inverse
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ʰ`:std.prec.max_plus := inverse
variables {ob : Type} [C : precategory ob]
variables {a b c : ob} {g : b ⟶ c} {f : a ⟶ b} {h : b ⟶ a}
include C
inductive is_section [class] (f : a ⟶ b) : Type
:= mk : ∀{g}, g ∘ f = id → is_section f
inductive is_retraction [class] (f : a ⟶ b) : Type
:= mk : ∀{g}, f ∘ g = id → is_retraction f
inductive is_iso [class] (f : a ⟶ b) : Type
:= mk : ∀{g}, g ∘ f = id → f ∘ g = id → is_iso f
attribute is_iso [multiple-instances]
definition retraction_of (f : a ⟶ b) [H : is_section f] : hom b a :=
is_section.rec (λg h, g) H
definition section_of (f : a ⟶ b) [H : is_retraction f] : hom b a :=
is_retraction.rec (λg h, g) H
definition inverse (f : a ⟶ b) [H : is_iso f] : hom b a :=
is_iso.rec (λg h1 h2, g) H
postfix `⁻¹` := inverse
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ʰ`:100 := inverse
theorem inverse_compose (f : a ⟶ b) [H : is_iso f] : f⁻¹ ∘ f = id :=
is_iso.rec (λg h1 h2, h1) H
theorem compose_inverse (f : a ⟶ b) [H : is_iso f] : f ∘ f⁻¹ = id :=
is_iso.rec (λg h1 h2, h2) H
theorem retraction_compose (f : a ⟶ b) [H : is_section f] : retraction_of f ∘ f = id :=
is_section.rec (λg h, h) H
theorem compose_section (f : a ⟶ b) [H : is_retraction f] : f ∘ section_of f = id :=
is_retraction.rec (λg h, h) H
theorem is_section_of_is_iso [instance] (f : a ⟶ b) [H : is_iso f] : is_section f :=
definition iso_imp_retraction [instance] [priority 300] [reducible]
(f : a ⟶ b) [H : is_iso f] : is_section f :=
is_section.mk !inverse_compose
theorem is_retraction_of_is_iso [instance] (f : a ⟶ b) [H : is_iso f] : is_retraction f :=
definition iso_imp_section [instance] [priority 300] [reducible]
(f : a ⟶ b) [H : is_iso f] : is_retraction f :=
is_retraction.mk !compose_inverse
theorem is_iso_id [instance] : is_iso (ID a) :=
definition is_iso_id [instance] [priority 500] (a : ob) : is_iso (ID a) :=
is_iso.mk !id_compose !id_compose
theorem is_iso_inverse [instance] (f : a ⟶ b) [H : is_iso f] : is_iso (f⁻¹) :=
definition is_iso_inverse [instance] [priority 200] (f : a ⟶ b) [H : is_iso f] : is_iso f⁻¹ :=
is_iso.mk !compose_inverse !inverse_compose
theorem left_inverse_eq_right_inverse {f : a ⟶ b} {g g' : hom b a}
definition left_inverse_eq_right_inverse {f : a ⟶ b} {g g' : hom b a}
(Hl : g ∘ f = id) (Hr : f ∘ g' = id) : g = g' :=
by rewrite [-(id_right g), -Hr, assoc, Hl, id_left]
theorem retraction_eq_intro [H : is_section f] (H2 : f ∘ h = id) : retraction_of f = h :=
definition retraction_eq_intro [H : is_section f] (H2 : f ∘ h = id) : retraction_of f = h :=
left_inverse_eq_right_inverse !retraction_compose H2
theorem section_eq_intro [H : is_retraction f] (H2 : h ∘ f = id) : section_of f = h :=
definition section_eq_intro [H : is_retraction f] (H2 : h ∘ f = id) : section_of f = h :=
(left_inverse_eq_right_inverse H2 !compose_section)⁻¹
theorem inverse_eq_intro_right [H : is_iso f] (H2 : f ∘ h = id) : f⁻¹ = h :=
definition inverse_eq_intro_right [H : is_iso f] (H2 : f ∘ h = id) : f⁻¹ = h :=
left_inverse_eq_right_inverse !inverse_compose H2
theorem inverse_eq_intro_left [H : is_iso f] (H2 : h ∘ f = id) : f⁻¹ = h :=
definition inverse_eq_intro_left [H : is_iso f] (H2 : h ∘ f = id) : f⁻¹ = h :=
(left_inverse_eq_right_inverse H2 !compose_inverse)⁻¹
theorem section_of_eq_retraction_of (f : a ⟶ b) [Hl : is_section f] [Hr : is_retraction f] :
definition section_eq_retraction (f : a ⟶ b) [Hl : is_section f] [Hr : is_retraction f] :
retraction_of f = section_of f :=
retraction_eq_intro !compose_section
theorem is_iso_of_is_retraction_of_is_section (f : a ⟶ b) [Hl : is_section f] [Hr : is_retraction f]
: is_iso f :=
is_iso.mk ((section_of_eq_retraction_of f) ▹ (retraction_compose f)) (compose_section f)
definition section_retraction_imp_iso (f : a ⟶ b) [Hl : is_section f] [Hr : is_retraction f]
: is_iso f :=
is_iso.mk ((section_eq_retraction f) ▹ (retraction_compose f)) (compose_section f)
theorem inverse_unique (H H' : is_iso f) : @inverse _ _ _ _ f H = @inverse _ _ _ _ f H' :=
definition inverse_unique (H H' : is_iso f) : @inverse _ _ _ _ f H = @inverse _ _ _ _ f H' :=
inverse_eq_intro_left !inverse_compose
theorem inverse_involutive (f : a ⟶ b) [H1 : is_iso f] [H2 : is_iso (f⁻¹)] : (f⁻¹)⁻¹ = f :=
definition inverse_involutive (f : a ⟶ b) [H : is_iso f] : (f⁻¹)⁻¹ = f :=
inverse_eq_intro_right !inverse_compose
theorem retraction_of_id : retraction_of (ID a) = id :=
definition retraction_of_id (a : ob) : retraction_of (ID a) = id :=
retraction_eq_intro !id_compose
theorem section_of_id : section_of (ID a) = id :=
definition section_of_id (a : ob) : section_of (ID a) = id :=
section_eq_intro !id_compose
theorem id_inverse [H : is_iso (ID a)] : (ID a)⁻¹ = id :=
definition iso_of_id (a : ob) [H : is_iso (ID a)] : (ID a)⁻¹ = id :=
inverse_eq_intro_left !id_compose
theorem is_section_comp [instance] [Hf : is_section f] [Hg : is_section g]
: is_section (g ∘ f) :=
definition composition_is_section [instance] [priority 150] (g : b ⟶ c) (f : a ⟶ b)
[Hf : is_section f] [Hg : is_section g] : is_section (g ∘ f) :=
is_section.mk
(show (retraction_of f ∘ retraction_of g) ∘ g ∘ f = id,
by rewrite [-assoc, assoc _ g f, retraction_compose, id_left, retraction_compose])
theorem is_retraction_comp [instance] [Hf : is_retraction f] [Hg : is_retraction g]
: is_retraction (g ∘ f) :=
definition composition_is_retraction [instance] [priority 150] (g : b ⟶ c) (f : a ⟶ b)
[Hf : is_retraction f] [Hg : is_retraction g] : is_retraction (g ∘ f) :=
is_retraction.mk
(show (g ∘ f) ∘ section_of f ∘ section_of g = id,
by rewrite [-assoc, {f ∘ _}assoc, compose_section, id_left, compose_section])
theorem is_inverse_comp [instance] [Hf : is_iso f] [Hg : is_iso g] : is_iso (g ∘ f) :=
!is_iso_of_is_retraction_of_is_section
definition composition_is_inverse [instance] [priority 150] (g : b ⟶ c) (f : a ⟶ b)
[Hf : is_iso f] [Hg : is_iso g] : is_iso (g ∘ f) :=
!section_retraction_imp_iso
structure isomorphic (a b : ob) :=
(iso : hom a b)
[is_iso : is_iso iso]
[struct : is_iso iso]
infix `≅`:50 := morphism.isomorphic
attribute isomorphic.is_iso [instance]
attribute isomorphic.struct [instance] [priority 400]
namespace isomorphic
attribute to_fun [coercion]
definition refl (a : ob) : a ≅ a :=
mk id
protected definition refl (a : ob) : a ≅ a :=
mk (ID a)
definition symm ⦃a b : ob⦄ (H : a ≅ b) : b ≅ a :=
mk (inverse (iso H))
protected definition symm ⦃a b : ob⦄ (H : a ≅ b) : b ≅ a :=
mk (iso H)⁻¹
definition trans ⦃a b c : ob⦄ (H1 : a ≅ b) (H2 : b ≅ c) : a ≅ c :=
protected definition trans ⦃a b c : ob⦄ (H1 : a ≅ b) (H2 : b ≅ c) : a ≅ c :=
mk (iso H2 ∘ iso H1)
end isomorphic
inductive is_mono [class] (f : a ⟶ b) : Type :=
mk : (∀c (g h : hom c a), f ∘ g = f ∘ h → g = h) → is_mono f
inductive is_epi [class] (f : a ⟶ b) : Type :=
mk : (∀c (g h : hom b c), g ∘ f = h ∘ f → g = h) → is_epi f
structure is_mono [class] (f : a ⟶ b) :=
(elim : ∀c (g h : hom c a), f ∘ g = f ∘ h → g = h)
structure is_epi [class] (f : a ⟶ b) :=
(elim : ∀c (g h : hom b c), g ∘ f = h ∘ f → g = h)
theorem is_mono.elim [H : is_mono f] {g h : c ⟶ a} (H2 : f ∘ g = f ∘ h) : g = h
:= is_mono.rec (λH3, H3 c g h H2) H
theorem is_epi.elim [H : is_epi f] {g h : b ⟶ c} (H2 : g ∘ f = h ∘ f) : g = h
:= is_epi.rec (λH3, H3 c g h H2) H
theorem is_mono_of_is_section [instance] (f : a ⟶ b) [H : is_section f] : is_mono f :=
definition is_mono_of_is_section [instance] (f : a ⟶ b) [H : is_section f] : is_mono f :=
is_mono.mk
(λ c g h H,
calc
@ -153,7 +144,7 @@ namespace morphism
... = id ∘ h : by rewrite retraction_compose
... = h : by rewrite id_left)
theorem is_epi_of_is_retraction [instance] (f : a ⟶ b) [H : is_retraction f] : is_epi f :=
definition is_epi_of_is_retraction [instance] (f : a ⟶ b) [H : is_retraction f] : is_epi f :=
is_epi.mk
(λ c g h H,
calc
@ -163,72 +154,66 @@ namespace morphism
... = h ∘ id : by rewrite compose_section
... = h : by rewrite id_right)
theorem is_mono_comp [instance] [Hf : is_mono f] [Hg : is_mono g] : is_mono (g ∘ f) :=
definition is_mono_comp [instance] (g : b ⟶ c) (f : a ⟶ b) [Hf : is_mono f] [Hg : is_mono g]
: is_mono (g ∘ f) :=
is_mono.mk
(λ d h₁ h₂ H,
have H2 : g ∘ (f ∘ h₁) = g ∘ (f ∘ h₂),
begin
rewrite *assoc, exact H
end,
is_mono.elim (is_mono.elim H2))
!is_mono.elim (!is_mono.elim H2))
theorem is_epi_comp [instance] [Hf : is_epi f] [Hg : is_epi g] : is_epi (g ∘ f) :=
definition is_epi_comp [instance] (g : b ⟶ c) (f : a ⟶ b) [Hf : is_epi f] [Hg : is_epi g]
: is_epi (g ∘ f) :=
is_epi.mk
(λ d h₁ h₂ H,
have H2 : (h₁ ∘ g) ∘ f = (h₂ ∘ g) ∘ f,
begin
rewrite -*assoc, exact H
end,
is_epi.elim (is_epi.elim H2))
!is_epi.elim (!is_epi.elim H2))
end morphism
namespace morphism
--rewrite lemmas for inverses, modified from
--https://github.com/JasonGross/HoTT-categories/blob/master/theories/Categories/Category/Morphisms.v
/-
rewrite lemmas for inverses, modified from
https://github.com/JasonGross/HoTT-categories/blob/master/theories/Categories/Category/Morphisms.v
-/
namespace iso
section
variables {ob : Type} [C : precategory ob] include C
variables {a b c d : ob} (f : b ⟶ a)
variables {a b c d : ob} (f : b ⟶ a)
(r : c ⟶ d) (q : b ⟶ c) (p : a ⟶ b)
(g : d ⟶ c)
variable [Hq : is_iso q] include Hq
theorem compose_pV : q ∘ q⁻¹ = id := !compose_inverse
theorem compose_Vp : q⁻¹ ∘ q = id := !inverse_compose
theorem compose_V_pp : q⁻¹ ∘ (q ∘ p) = p :=
by rewrite [assoc, inverse_compose, id_left]
theorem compose_p_Vp : q ∘ (q⁻¹ ∘ g) = g :=
by rewrite [assoc, compose_inverse, id_left]
theorem compose_pp_V : (r ∘ q) ∘ q⁻¹ = r :=
definition compose_pV : q ∘ q⁻¹ = id := !compose_inverse
definition compose_Vp : q⁻¹ ∘ q = id := !inverse_compose
definition compose_V_pp : q⁻¹ ∘ (q ∘ p) = p :=
by rewrite [assoc, inverse_compose, id_left]
definition compose_p_Vp : q ∘ (q⁻¹ ∘ g) = g :=
by rewrite [assoc, compose_inverse, id_left]
definition compose_pp_V : (r ∘ q) ∘ q⁻¹ = r :=
by rewrite [-assoc, compose_inverse, id_right]
theorem compose_pV_p : (f ∘ q⁻¹) ∘ q = f :=
definition compose_pV_p : (f ∘ q⁻¹) ∘ q = f :=
by rewrite [-assoc, inverse_compose, id_right]
theorem con_inv [H' : is_iso p] [Hpq : is_iso (q ∘ p)] : (q ∘ p)⁻¹ = p⁻¹ ∘ q⁻¹ :=
definition con_inv [Hp : is_iso p] [Hpq : is_iso (q ∘ p)] : (q ∘ p)⁻¹ʰ = p⁻¹ʰ ∘ q⁻¹ʰ :=
inverse_eq_intro_left
(show (p⁻¹ ∘ (q⁻¹)) ∘ q ∘ p = id, from
(show (p⁻¹ʰ ∘ q⁻¹ʰ) ∘ q ∘ p = id, from
by rewrite [-assoc, compose_V_pp, inverse_compose])
--the proof using calc is hard for the unifier (needs ~90k steps)
-- inverse_eq_intro_left
-- (calc
-- (p⁻¹ ∘ (q⁻¹)) ∘ q ∘ p = p⁻¹ ∘ (q⁻¹ ∘ (q ∘ p)) : assoc (p⁻¹) (q⁻¹) (q ∘ p)⁻¹
-- ... = (p⁻¹) ∘ p : congr_arg (λx, p⁻¹ ∘ x) (compose_V_pp q p)
-- ... = id : inverse_compose p)
theorem inv_con_inv_left [H' : is_iso g] : (q⁻¹ ∘ g)⁻¹ = g⁻¹ ∘ q :=
inverse_involutive q ▹ con_inv (q⁻¹) g
definition inv_con_inv_left [H' : is_iso g] : (q⁻¹ ∘ g)⁻¹ = g⁻¹ ∘ q :=
inverse_involutive q ▹ con_inv q⁻¹ g
theorem inv_con_inv_right [H' : is_iso f] : (q ∘ f⁻¹)⁻¹ = f ∘ q⁻¹ :=
inverse_involutive f ▹ con_inv q (f⁻¹)
theorem inv_con_inv_inv [H' : is_iso r] : (q⁻¹ ∘ r⁻¹)⁻¹ = r ∘ q :=
inverse_involutive r ▹ inv_con_inv_left q (r⁻¹)
definition inv_con_inv_right [H' : is_iso f] : (q ∘ f⁻¹)⁻¹ = f ∘ q⁻¹ :=
inverse_involutive f ▹ con_inv q f⁻¹
definition inv_con_inv_inv [H' : is_iso r] : (q⁻¹ ∘ r⁻¹)⁻¹ = r ∘ q :=
inverse_involutive r ▹ inv_con_inv_left q r⁻¹
end
section
variables {ob : Type} {C : precategory ob} include C
variables {d c b a : ob}
@ -239,22 +224,24 @@ namespace morphism
{y : d ⟶ b} {w : c ⟶ a}
variable [Hq : is_iso q] include Hq
theorem con_eq_of_eq_inv_con (H : y = q⁻¹ ∘ g) : q ∘ y = g := H⁻¹ ▹ compose_p_Vp q g
theorem con_eq_of_eq_con_inv (H : w = f ∘ q⁻¹) : w ∘ q = f := H⁻¹ ▹ compose_pV_p f q
theorem inv_con_eq_of_eq_con (H : z = q ∘ p) : q⁻¹ ∘ z = p := H⁻¹ ▹ compose_V_pp q p
theorem con_inv_eq_of_eq_con (H : x = r ∘ q) : x ∘ q⁻¹ = r := H⁻¹ ▹ compose_pp_V r q
theorem eq_con_of_inv_con_eq (H : q⁻¹ ∘ g = y) : g = q ∘ y := con_eq_of_eq_inv_con (H⁻¹)⁻¹
theorem eq_con_of_con_inv_eq (H : f ∘ q⁻¹ = w) : f = w ∘ q := con_eq_of_eq_con_inv (H⁻¹)⁻¹
theorem eq_inv_con_of_con_eq (H : q ∘ p = z) : p = q⁻¹ ∘ z := inv_con_eq_of_eq_con (H⁻¹)⁻¹
theorem eq_con_inv_of_con_eq (H : r ∘ q = x) : r = x ∘ q⁻¹ := con_inv_eq_of_eq_con (H⁻¹)⁻¹
theorem eq_inv_of_con_eq_idp' (H : h ∘ q = id) : h = q⁻¹ := inverse_eq_intro_left H⁻¹
theorem eq_inv_of_con_eq_idp (H : q ∘ h = id) : h = q⁻¹ := inverse_eq_intro_right H⁻¹
theorem eq_of_con_inv_eq_idp (H : i ∘ q⁻¹ = id) : i = q := eq_inv_of_con_eq_idp' H ⬝ inverse_involutive q
theorem eq_of_inv_con_eq_idp (H : q⁻¹ ∘ i = id) : i = q := eq_inv_of_con_eq_idp H ⬝ inverse_involutive q
theorem eq_of_idp_eq_con_inv (H : id = i ∘ q⁻¹) : q = i := eq_of_con_inv_eq_idp (H⁻¹)⁻¹
theorem eq_of_idp_eq_inv_con (H : id = q⁻¹ ∘ i) : q = i := eq_of_inv_con_eq_idp (H⁻¹)⁻¹
theorem inv_eq_of_idp_eq_con (H : id = h ∘ q) : q⁻¹ = h := eq_inv_of_con_eq_idp' (H⁻¹)⁻¹
theorem inv_eq_of_idp_eq_con' (H : id = q ∘ h) : q⁻¹ = h := eq_inv_of_con_eq_idp (H⁻¹)⁻¹
definition con_eq_of_eq_inv_con (H : y = q⁻¹ ∘ g) : q ∘ y = g := H⁻¹ ▹ compose_p_Vp q g
definition con_eq_of_eq_con_inv (H : w = f ∘ q⁻¹) : w ∘ q = f := H⁻¹ ▹ compose_pV_p f q
definition inv_con_eq_of_eq_con (H : z = q ∘ p) : q⁻¹ ∘ z = p := H⁻¹ ▹ compose_V_pp q p
definition con_inv_eq_of_eq_con (H : x = r ∘ q) : x ∘ q⁻¹ = r := H⁻¹ ▹ compose_pp_V r q
definition eq_con_of_inv_con_eq (H : q⁻¹ ∘ g = y) : g = q ∘ y := (con_eq_of_eq_inv_con H⁻¹)⁻¹
definition eq_con_of_con_inv_eq (H : f ∘ q⁻¹ = w) : f = w ∘ q := (con_eq_of_eq_con_inv H⁻¹)⁻¹
definition eq_inv_con_of_con_eq (H : q ∘ p = z) : p = q⁻¹ ∘ z := (inv_con_eq_of_eq_con H⁻¹)⁻¹
definition eq_con_inv_of_con_eq (H : r ∘ q = x) : r = x ∘ q⁻¹ := (con_inv_eq_of_eq_con H⁻¹)⁻¹
definition eq_inv_of_con_eq_idp' (H : h ∘ q = id) : h = q⁻¹ := (inverse_eq_intro_left H)⁻¹
definition eq_inv_of_con_eq_idp (H : q ∘ h = id) : h = q⁻¹ := (inverse_eq_intro_right H)⁻¹
definition eq_of_con_inv_eq_idp (H : i ∘ q⁻¹ = id) : i = q :=
eq_inv_of_con_eq_idp' H ⬝ inverse_involutive q
definition eq_of_inv_con_eq_idp (H : q⁻¹ ∘ i = id) : i = q :=
eq_inv_of_con_eq_idp H ⬝ inverse_involutive q
definition eq_of_idp_eq_con_inv (H : id = i ∘ q⁻¹) : q = i := (eq_of_con_inv_eq_idp H⁻¹)⁻¹
definition eq_of_idp_eq_inv_con (H : id = q⁻¹ ∘ i) : q = i := (eq_of_inv_con_eq_idp H⁻¹)⁻¹
definition inv_eq_of_idp_eq_con (H : id = h ∘ q) : q⁻¹ = h := (eq_inv_of_con_eq_idp' H⁻¹)⁻¹
definition inv_eq_of_idp_eq_con' (H : id = q ∘ h) : q⁻¹ = h := (eq_inv_of_con_eq_idp H⁻¹)⁻¹
end
end iso

View file

@ -3,7 +3,7 @@ Copyright (c) 2014 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Module: algebra.precategory.yoneda
Author: Floris van Doorn
Authors: Floris van Doorn
-/
--note: modify definition in category.set

View file

@ -127,7 +127,7 @@ theorem weak_funext_of_ua : weak_funext :=
have tU : is_contr (Π x, U x),
from tU',
have tlast : is_contr (Πx, P x),
from eq.transport _ (p⁻¹) tU,
from eq.transport _ p⁻¹ tU,
tlast
)

View file

@ -3,7 +3,7 @@ Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Module: init.bool
Author: Leonardo de Moura
Authors: Leonardo de Moura
-/
prelude

View file

@ -54,10 +54,22 @@ inductive pos_num : Type :=
| bit1 : pos_num → pos_num
| bit0 : pos_num → pos_num
namespace pos_num
definition succ (a : pos_num) : pos_num :=
pos_num.rec_on a (bit0 one) (λn r, bit0 r) (λn r, bit1 n)
end pos_num
inductive num : Type :=
| zero : num
| pos : pos_num → num
namespace num
open pos_num
definition succ (a : num) : num :=
num.rec_on a (pos one) (λp, pos (succ p))
end num
inductive bool : Type :=
| ff : bool
| tt : bool

View file

@ -33,7 +33,7 @@ namespace is_equiv
/- Some instances and closure properties of equivalences -/
postfix `⁻¹` := inv
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ᵉ`:100 := inv
postfix [parsing-only] `⁻¹ᵉ`:std.prec.max_plus := inv
section
variables {A B C : Type} (f : A → B) (g : B → C) {f' : A → B}
@ -48,7 +48,7 @@ namespace is_equiv
(λa, ap (inv f) (sect g (f a)) ⬝ sect f a)
(λa, (whisker_left _ (adj g (f a))) ⬝
(ap_con g _ _)⁻¹ ⬝
ap02 g (ap_con_eq_con (retr f) (sect g (f a))⁻¹ ⬝
ap02 g ((ap_con_eq_con (retr f) (sect g (f a)))⁻¹ ⬝
(ap_compose (inv f) f _ ◾ adj f a) ⬝
(ap_con f _ _)⁻¹
) ⬝
@ -76,18 +76,18 @@ namespace is_equiv
... = ap f (ap invf ff'a) ⬝ retrf'a : ap_compose invf f,
have eq2 : _ = _,
from calc retrf'a
= (ap f (ap invf ff'a))⁻¹ ⬝ (ap f secta ⬝ ff'a) : eq_inv_con_of_con_eq _ _ _ (eq1⁻¹)
... = ap f (ap invf ff'a)⁻¹ ⬝ (ap f secta ⬝ Hty a) : ap_inv invf ff'a
... = ap f (ap invf ff'a)⁻¹ ⬝ (Hty (invf (f a)) ⬝ ap f' secta) : ap_con_eq_con_ap
... = (ap f (ap invf ff'a)⁻¹ ⬝ Hty (invf (f a))) ⬝ ap f' secta : con.assoc
... = (ap f ((ap invf ff'a)⁻¹) ⬝ Hty (invf (f a))) ⬝ ap f' secta : ap_inv
... = (Hty (invf (f' a)) ⬝ ap f' ((ap invf ff'a)⁻¹)) ⬝ ap f' secta : ap_con_eq_con_ap
= (ap f (ap invf ff'a))⁻¹ ⬝ (ap f secta ⬝ ff'a) : eq_inv_con_of_con_eq _ _ _ eq1⁻¹
... = (ap f (ap invf ff'a))⁻¹ ⬝ (ap f secta ⬝ Hty a) : ap_inv invf ff'a
... = (ap f (ap invf ff'a))⁻¹ ⬝ (Hty (invf (f a)) ⬝ ap f' secta) : ap_con_eq_con_ap
... = ((ap f (ap invf ff'a))⁻¹ ⬝ Hty (invf (f a))) ⬝ ap f' secta : con.assoc
... = (ap f (ap invf ff'a)⁻¹ ⬝ Hty (invf (f a))) ⬝ ap f' secta : ap_inv
... = (Hty (invf (f' a)) ⬝ ap f' (ap invf ff'a)⁻¹) ⬝ ap f' secta : ap_con_eq_con_ap
... = (Hty (invf (f' a)) ⬝ (ap f' (ap invf ff'a))⁻¹) ⬝ ap f' secta : ap_inv
... = Hty (invf (f' a)) ⬝ ((ap f' (ap invf ff'a))⁻¹ ⬝ ap f' secta) : con.assoc,
have eq3 : _ = _,
from calc (Hty (invf (f' a)))⁻¹ ⬝ retrf'a
= (ap f' (ap invf ff'a))⁻¹ ⬝ ap f' secta : inv_con_eq_of_eq_con _ _ _ eq2
... = (ap f' ((ap invf ff'a)⁻¹)) ⬝ ap f' secta : ap_inv
... = (ap f' (ap invf ff'a)⁻¹) ⬝ ap f' secta : ap_inv
... = ap f' ((ap invf ff'a)⁻¹ ⬝ secta) : ap_con,
eq3) in
is_equiv.mk (inv f) sect' retr' adj'
@ -103,31 +103,31 @@ namespace is_equiv
private definition adjointify_adj' : Π (x : A), ret (f x) = ap f (adjointify_sect' x) :=
(λ (a : A),
let fgretrfa := ap f (ap g (ret (f a))) in
let fgfinvsect := ap f (ap g (ap f ((sec a)⁻¹))) in
let fgfinvsect := ap f (ap g (ap f (sec a)⁻¹)) in
let fgfa := f (g (f a)) in
let retrfa := ret (f a) in
have eq1 : ap f (sec a) = _,
from calc ap f (sec a)
= idp ⬝ ap f (sec a) : !idp_con⁻¹
... = (ret (f a) ⬝ (ret (f a)⁻¹)) ⬝ ap f (sec a) : {!con.left_inv⁻¹}
... = ((ret (fgfa))⁻¹ ⬝ ap (f ∘ g) (ret (f a))) ⬝ ap f (sec a) : {!con_ap_eq_con⁻¹}
... = ((ret (fgfa))⁻¹ ⬝ fgretrfa) ⬝ ap f (sec a) : {ap_compose g f _}
... = (ret (fgfa))⁻¹ ⬝ (fgretrfa ⬝ ap f (sec a)) : !con.assoc,
... = (ret (f a) ⬝ (ret (f a))⁻¹) ⬝ ap f (sec a) : {!con.left_inv⁻¹}
... = ((ret fgfa)⁻¹ ⬝ ap (f ∘ g) (ret (f a))) ⬝ ap f (sec a) : {!con_ap_eq_con⁻¹}
... = ((ret fgfa)⁻¹ ⬝ fgretrfa) ⬝ ap f (sec a) : {ap_compose g f _}
... = (ret fgfa)⁻¹ ⬝ (fgretrfa ⬝ ap f (sec a)) : !con.assoc,
have eq2 : ap f (sec a) ⬝ idp = (ret fgfa)⁻¹ ⬝ (fgretrfa ⬝ ap f (sec a)),
from !con_idp ⬝ eq1,
have eq3 : idp = _,
from calc idp
= (ap f (sec a))⁻¹ ⬝ ((ret fgfa)⁻¹ ⬝ (fgretrfa ⬝ ap f (sec a))) : eq_inv_con_of_con_eq _ _ _ eq2
... = (ap f (sec a)⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : !con.assoc'
... = (ap f ((sec a)⁻¹) ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : {!ap_inv⁻¹}
... = ((ap f ((sec a)⁻¹) ⬝ (ret fgfa)⁻¹) ⬝ fgretrfa) ⬝ ap f (sec a) : !con.assoc'
... = ((retrfa⁻¹ ⬝ ap (f ∘ g) (ap f ((sec a)⁻¹))) ⬝ fgretrfa) ⬝ ap f (sec a) : {!con_ap_eq_con⁻¹}
... = ((ap f (sec a))⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : !con.assoc'
... = (ap f (sec a)⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : {!ap_inv⁻¹}
... = ((ap f (sec a)⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ fgretrfa) ⬝ ap f (sec a) : !con.assoc'
... = ((retrfa⁻¹ ⬝ ap (f ∘ g) (ap f (sec a)⁻¹)) ⬝ fgretrfa) ⬝ ap f (sec a) : {!con_ap_eq_con⁻¹}
... = ((retrfa⁻¹ ⬝ fgfinvsect) ⬝ fgretrfa) ⬝ ap f (sec a) : {ap_compose g f _}
... = (retrfa⁻¹ ⬝ (fgfinvsect ⬝ fgretrfa)) ⬝ ap f (sec a) : {!con.assoc'⁻¹}
... = retrfa⁻¹ ⬝ ap f (ap g (ap f ((sec a)⁻¹)) ⬝ ap g (ret (f a))) ⬝ ap f (sec a) : {!ap_con⁻¹}
... = retrfa⁻¹ ⬝ (ap f (ap g (ap f ((sec a)⁻¹)) ⬝ ap g (ret (f a))) ⬝ ap f (sec a)) : !con.assoc'⁻¹
... = retrfa⁻¹ ⬝ ap f ((ap g (ap f ((sec a)⁻¹)) ⬝ ap g (ret (f a))) ⬝ sec a) : {!ap_con⁻¹},
have eq4 : ret (f a) = ap f ((ap g (ap f ((sec a)⁻¹)) ⬝ ap g (ret (f a))) ⬝ sec a),
... = retrfa⁻¹ ⬝ ap f (ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ ap f (sec a) : {!ap_con⁻¹}
... = retrfa⁻¹ ⬝ (ap f (ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ ap f (sec a)) : !con.assoc'⁻¹
... = retrfa⁻¹ ⬝ ap f ((ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ sec a) : {!ap_con⁻¹},
have eq4 : ret (f a) = ap f ((ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ sec a),
from eq_of_idp_eq_inv_con _ _ eq3,
eq4)
@ -151,38 +151,38 @@ namespace is_equiv
variable (g : B → C)
definition cancel_right (g : B → C) [Hgf : is_equiv (g ∘ f)] : (is_equiv g) :=
have Hfinv [visible] : is_equiv (f⁻¹), from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose (f⁻¹) (g ∘ f)) (λb, ap g (@retr _ _ f _ b))
have Hfinv [visible] : is_equiv f⁻¹, from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose f⁻¹ (g ∘ f)) (λb, ap g (@retr _ _ f _ b))
definition cancel_left (g : C → A) [Hgf : is_equiv (f ∘ g)] : (is_equiv g) :=
have Hfinv [visible] : is_equiv (f⁻¹), from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose (f ∘ g) (f⁻¹)) (λa, sect f (g a))
have Hfinv [visible] : is_equiv f⁻¹, from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose (f ∘ g) f⁻¹) (λa, sect f (g a))
--Rewrite rules
definition eq_of_eq_inv {x : A} {y : B} (p : x = (inv f) y) : (f x = y) :=
(ap f p) ⬝ (@retr _ _ f _ y)
definition eq_of_inv_eq {x : A} {y : B} (p : (inv f) y = x) : (y = f x) :=
(eq_of_eq_inv f (p⁻¹))⁻¹
(eq_of_eq_inv f p⁻¹)⁻¹
definition inv_eq_of_eq {x : B} {y : A} (p : x = f y) : (inv f) x = y :=
ap (f⁻¹) p ⬝ sect f y
ap f⁻¹ p ⬝ sect f y
definition eq_inv_of_eq {x : B} {y : A} (p : f y = x) : y = (inv f) x :=
(inv_eq_of_eq f (p⁻¹))⁻¹
(inv_eq_of_eq f p⁻¹)⁻¹
definition is_equiv_ap [instance] (x y : A) : is_equiv (ap f) :=
adjointify (ap f)
(λq, (inverse (sect f x)) ⬝ ap (f⁻¹) q ⬝ sect f y)
(λq, (inverse (sect f x)) ⬝ ap f⁻¹ q ⬝ sect f y)
(λq, !ap_con
⬝ whisker_right !ap_con _
⬝ ((!ap_inv ⬝ inverse2 ((adj f _)⁻¹))
◾ (inverse (ap_compose (f⁻¹) f _))
⬝ ((!ap_inv ⬝ inverse2 (adj f _)⁻¹)
◾ (inverse (ap_compose f⁻¹ f _))
◾ (adj f _)⁻¹)
⬝ con_ap_con_eq_con_con (retr f) _ _
⬝ whisker_right !con.right_inv _
⬝ !idp_con)
(λp, whisker_right (whisker_left _ ((ap_compose f (f⁻¹) _)⁻¹)) _
(λp, whisker_right (whisker_left _ (ap_compose f f⁻¹ _)⁻¹) _
⬝ con_ap_con_eq_con_con (sect f) _ _
⬝ whisker_right !con.right_inv _
⬝ !idp_con)
@ -211,7 +211,7 @@ namespace is_equiv
--Transporting is an equivalence
definition is_equiv_tr [instance] {A : Type} (P : A → Type) {x y : A} (p : x = y) : (is_equiv (transport P p)) :=
is_equiv.mk (transport P (p⁻¹)) (tr_inv_tr P p) (inv_tr_tr P p) (tr_inv_tr_lemma P p)
is_equiv.mk (transport P p⁻¹) (tr_inv_tr P p) (inv_tr_tr P p) (tr_inv_tr_lemma P p)
end is_equiv
@ -237,7 +237,7 @@ namespace equiv
equiv.mk id !is_equiv_id
protected definition symm (f : A ≃ B) : B ≃ A :=
equiv.mk (f⁻¹) !is_equiv_inv
equiv.mk f⁻¹ !is_equiv_inv
protected definition trans (f : A ≃ B) (g: B ≃ C) : A ≃ C :=
equiv.mk (g ∘ f) !is_equiv_compose

View file

@ -94,7 +94,7 @@ namespace ne
assume H, H rfl
definition symm : a ≠ b → b ≠ a :=
assume (H : a ≠ b) (H₁ : b = a), H (H₁⁻¹)
assume (H : a ≠ b) (H₁ : b = a), H H₁⁻¹
end ne
section

View file

@ -2,7 +2,7 @@
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Module init.num
Module: init.num
Authors: Leonardo de Moura
-/
@ -14,9 +14,6 @@ definition pos_num.is_inhabited [instance] : inhabited pos_num :=
inhabited.mk pos_num.one
namespace pos_num
definition succ (a : pos_num) : pos_num :=
pos_num.rec_on a (bit0 one) (λn r, bit0 r) (λn r, bit1 n)
definition is_one (a : pos_num) : bool :=
pos_num.rec_on a tt (λn r, ff) (λn r, ff)
@ -56,9 +53,6 @@ inhabited.mk num.zero
namespace num
open pos_num
definition succ (a : num) : num :=
num.rec_on a (pos one) (λp, pos (succ p))
definition pred (a : num) : num :=
num.rec_on a zero (λp, cond (is_one p) zero (pos (pred p)))

View file

@ -43,7 +43,7 @@ namespace eq
notation p₁ ⬝ p₂ := concat p₁ p₂
notation p ⁻¹ := inverse p
--a second notation for the inverse, which is not overloaded
postfix [parsing-only] `⁻¹ᵖ`:100 := inverse
postfix [parsing-only] `⁻¹ᵖ`:std.prec.max_plus := inverse
/- The 1-dimensional groupoid structure -/
@ -272,10 +272,10 @@ namespace eq
eq.rec_on q (eq.rec_on p (take r, con.assoc _ _ _)) r
-- Functions commute with path inverses.
definition ap_inv' (f : A → B) {x y : A} (p : x = y) : (ap f p)⁻¹ = ap f (p⁻¹) :=
definition ap_inv' (f : A → B) {x y : A} (p : x = y) : (ap f p)⁻¹ = ap f p⁻¹ :=
eq.rec_on p idp
definition ap_inv {A B : Type} (f : A → B) {x y : A} (p : x = y) : ap f (p⁻¹) = (ap f p)⁻¹ :=
definition ap_inv {A B : Type} (f : A → B) {x y : A} (p : x = y) : ap f p⁻¹ = (ap f p)⁻¹ :=
eq.rec_on p idp
-- [ap] itself is functorial in the first argument.
@ -380,7 +380,7 @@ namespace eq
eq.rec_on h (take h', eq.rec_on h' idp) h'
definition apD10_inv {f g : Πx : A, P x} (h : f = g) (x : A) :
apD10 (h⁻¹) x = (apD10 h x)⁻¹ :=
apD10 h⁻¹ x = (apD10 h x)⁻¹ :=
eq.rec_on h idp
definition ap10_idp {f : A → B} (x : A) : ap10 (refl f) x = idp := idp
@ -388,7 +388,7 @@ namespace eq
definition ap10_con {f f' f'' : A → B} (h : f = f') (h' : f' = f'') (x : A) :
ap10 (h ⬝ h') x = ap10 h x ⬝ ap10 h' x := apD10_con h h' x
definition ap10_inv {f g : A → B} (h : f = g) (x : A) : ap10 (h⁻¹) x = (ap10 h x)⁻¹ :=
definition ap10_inv {f g : A → B} (h : f = g) (x : A) : ap10 h⁻¹ x = (ap10 h x)⁻¹ :=
apD10_inv h x
-- [ap10] also behaves nicely on paths produced by [ap]
@ -408,11 +408,11 @@ namespace eq
definition tr_inv_tr (P : A → Type) {x y : A} (p : x = y) (z : P y) :
p ▹ p⁻¹ ▹ z = z :=
(tr_con P (p⁻¹) p z)⁻¹ ⬝ ap (λr, transport P r z) (con.right_inv p)
(tr_con P p⁻¹ p z)⁻¹ ⬝ ap (λr, transport P r z) (con.right_inv p)
definition inv_tr_tr (P : A → Type) {x y : A} (p : x = y) (z : P x) :
p⁻¹ ▹ p ▹ z = z :=
(tr_con P p (p⁻¹) z)⁻¹ ⬝ ap (λr, transport P r z) (con.left_inv p)
(tr_con P p p⁻¹ z)⁻¹ ⬝ ap (λr, transport P r z) (con.left_inv p)
definition tr_con_lemma (P : A → Type)
{x y z w : A} (p : x = y) (q : y = z) (r : z = w) (u : P x) :
@ -454,7 +454,7 @@ namespace eq
eq.rec_on r1 (eq.rec_on r2 idp)
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)⁻¹) :=
transport2 Q r⁻¹ z = (transport2 Q r z)⁻¹ :=
eq.rec_on r idp
definition transportD2 (B C : A → Type) (D : Π(a:A), B a → C a → Type)
@ -640,7 +640,7 @@ namespace eq
apD02 f (r1 ⬝ r2) = apD02 f r1
⬝ whisker_left (transport2 P r1 (f x)) (apD02 f r2)
⬝ con.assoc' _ _ _
⬝ (whisker_right ((tr2_con P r1 r2 (f x))⁻¹) (apD f p3)) :=
⬝ (whisker_right (tr2_con P r1 r2 (f x))⁻¹ (apD f p3)) :=
eq.rec_on r2 (eq.rec_on r1 (eq.rec_on p1 idp))
end eq

View file

@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Module: init.reserved_notation
Authors: Leonardo de Moura
Basic datatypes
-/
prelude
@ -25,9 +23,21 @@ notation `take` binders `,` r:(scoped f, f) := r
/- Logical operations and relations -/
definition std.prec.max : num := 1024 -- reflects max precedence used internally
definition std.prec.max : num := 1024 -- the strength of application, identifiers, (, [, etc.
definition std.prec.arrow : num := 25
/-
The next definition is "max + 10". It can be used e.g. for postfix operations that should
be stronger than application.
-/
definition std.prec.max_plus :=
num.succ (num.succ (num.succ (num.succ (num.succ (num.succ (num.succ (num.succ (num.succ
(num.succ std.prec.max)))))))))
/- Logical operations and relations -/
reserve prefix `¬`:40
reserve prefix `~`:40
reserve infixr `∧`:35
@ -41,8 +51,8 @@ reserve infix `≠`:50
reserve infix `≈`:50
reserve infix ``:50
reserve infixr `∘`:60 -- input with \comp
reserve postfix `⁻¹`:100 --input with \sy or \-1 or \inv
reserve infixr `∘`:60 -- input with \comp
reserve postfix `⁻¹`:std.prec.max_plus -- input with \sy or \-1 or \inv
reserve infixl `⬝`:75
reserve infixr `▸`:75

View file

@ -102,11 +102,11 @@ namespace is_trunc
@contr_internal.contr A !is_trunc.to_internal a
definition center_eq [H : is_contr A] (x y : A) : x = y :=
contr x⁻¹ ⬝ (contr y)
(contr x)⁻¹ ⬝ (contr y)
definition hprop_eq {A : Type} [H : is_contr A] {x y : A} (p q : x = y) : p = q :=
have K : ∀ (r : x = y), center_eq x y = r, from (λ r, eq.rec_on r !con.right_inv),
K p⁻¹ ⬝ K q
(K p)⁻¹ ⬝ K q
definition is_contr_eq [instance] {A : Type} [H : is_contr A] (x y : A) : is_contr (x = y)
:=
@ -224,7 +224,7 @@ namespace is_trunc
trunc_index.rec_on n
(λA (HA : is_contr A) B f (H : is_equiv f), !is_contr_is_equiv_closed)
(λn IH A (HA : is_trunc n.+1 A) B f (H : is_equiv f), @is_trunc_succ_intro _ _ (λ x y : B,
IH (f⁻¹ x = f⁻¹ y) !is_trunc_eq (x = y) ((ap (f⁻¹))⁻¹) !is_equiv_inv))
IH (f⁻¹ x = f⁻¹ y) !is_trunc_eq (x = y) (ap f⁻¹)⁻¹ !is_equiv_inv))
A HA B f H
definition is_trunc_equiv_closed (n : trunc_index) (f : A ≃ B) [HA : is_trunc n A]
@ -252,7 +252,7 @@ namespace is_trunc
(is_equiv.mk (λ (u : unit), center A)
(λ (u : unit), unit.rec_on u idp)
(λ (x : A), contr x)
(λ (x : A), (!ap_constant)⁻¹))
(λ (x : A), !ap_constant⁻¹))
-- TODO: port "Truncated morphisms"

View file

@ -112,13 +112,13 @@ namespace path
definition isequiv_concat_l [instance] (p : a1 = a2) (a3 : A)
: is_equiv (@concat _ a1 a2 a3 p) :=
is_equiv.mk (concat (p⁻¹))
is_equiv.mk (concat p⁻¹)
(con_inv_cancel_left p)
(inv_con_cancel_left p)
(eq.rec_on p (λq, eq.rec_on q idp))
definition equiv_concat_l (p : a1 = a2) (a3 : A) : (a1 = a3) ≃ (a2 = a3) :=
equiv.mk (concat (p⁻¹)) _
equiv.mk (concat p⁻¹) _
definition isequiv_concat_r [instance] (p : a2 = a3) (a1 : A)
: is_equiv (λq : a1 = a2, q ⬝ p) :=

View file

@ -120,7 +120,7 @@ namespace pi
[H0 : is_equiv f0] [H1 : Πa', @is_equiv (B (f0 a')) (B' a') (f1 a')]
: is_equiv (pi_functor f0 f1) :=
begin
apply (adjointify (pi_functor f0 f1) (pi_functor (f0⁻¹)
apply (adjointify (pi_functor f0 f1) (pi_functor f0⁻¹
(λ(a : A) (b' : B' (f0⁻¹ a)), transport B (retr f0 a) ((f1 (f0⁻¹ a))⁻¹ b')))),
intro h, apply eq_of_homotopy,
unfold pi_functor, unfold function.compose, unfold function.id,

View file

@ -194,8 +194,8 @@ namespace sigma
definition is_equiv_sigma_functor [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)]
: is_equiv (sigma_functor f g) :=
adjointify (sigma_functor f g)
(sigma_functor (f⁻¹) (λ(a' : A') (b' : B' a'),
((g (f⁻¹ a'))⁻¹ (transport B' (retr f a'⁻¹) b'))))
(sigma_functor f⁻¹ (λ(a' : A') (b' : B' a'),
((g (f⁻¹ a'))⁻¹ (transport B' (retr f a')⁻¹ b'))))
begin
intro u',
cases u' with (a', b'),
@ -204,33 +204,25 @@ namespace sigma
-- end
-- "rewrite retr (g (f⁻¹ a'))"
apply concat, apply (ap (λx, (transport B' (retr f a') x))), apply (retr (g (f⁻¹ a'))),
show retr f a' ▹ (((retr f a') ⁻¹) ▹ b') = b',
show retr f a' ▹ ((retr f a')⁻¹ ▹ b') = b',
from tr_inv_tr B' (retr f a') b'
end
begin
intro u,
cases u with (a, b),
apply (sigma_eq (sect f a)),
show transport B (sect f a) (g (f⁻¹ (f a))⁻¹ (transport B' (retr f (f a)⁻¹) (g a b))) = b,
show transport B (sect f a) ((g (f⁻¹ (f a)))⁻¹ (transport B' (retr f (f a))⁻¹ (g a b))) = b,
from calc
transport B (sect f a) (g (f⁻¹ (f a))⁻¹ (transport B' (retr f (f a)⁻¹) (g a b)))
= g a⁻¹ (transport (B' ∘ f) (sect f a) (transport B' (retr f (f a)⁻¹) (g a b)))
: by rewrite (fn_tr_eq_tr_fn (sect f a) (λ a, g a⁻¹))
... = g a⁻¹ (transport B' (ap f (sect f a)) (transport B' (retr f (f a)⁻¹) (g a b)))
: ap (g a⁻¹) !transport_compose
... = g a⁻¹ (transport B' (ap f (sect f a)) (transport B' (ap f (sect f a)⁻¹) (g a b)))
: ap (λ x, g a⁻¹ (transport B' (ap f (sect f a)) (transport B' (x⁻¹) (g a b)))) (adj f a)
... = g a⁻¹ (g a b) : {!tr_inv_tr}
transport B (sect f a) ((g (f⁻¹ (f a)))⁻¹ (transport B' (retr f (f a))⁻¹ (g a b)))
= (g a)⁻¹ (transport (B' ∘ f) (sect f a) (transport B' (retr f (f a))⁻¹ (g a b)))
: by rewrite (fn_tr_eq_tr_fn (sect f a) (λ a, (g a)⁻¹))
... = (g a)⁻¹ (transport B' (ap f (sect f a)) (transport B' (retr f (f a))⁻¹ (g a b)))
: ap (g a)⁻¹ !transport_compose
... = (g a)⁻¹ (transport B' (ap f (sect f a)) (transport B' (ap f (sect f a))⁻¹ (g a b)))
: ap (λ x, (g a)⁻¹ (transport B' (ap f (sect f a)) (transport B' x⁻¹ (g a b)))) (adj f a)
... = (g a)⁻¹ (g a b) : {!tr_inv_tr}
... = b : by rewrite (sect (g a) b)
end
-- -- "rewrite fn_tr_eq_tr_fn"
-- apply concat, apply inverse, apply (fn_tr_eq_tr_fn (sect f a) (λ a, g a⁻¹)),
-- apply concat, apply (ap (g a⁻¹)),
-- -- "rewrite transport_compose"
-- apply concat, apply transport_compose,
-- -- "rewrite adj"
-- -- "rewrite tr_inv_tr"
-- apply sect,
definition sigma_equiv_sigma_of_is_equiv [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)]
: (Σa, B a) ≃ (Σa', B' a') :=
@ -250,13 +242,13 @@ namespace sigma
definition ap_sigma_functor_eq_dpair (p : a = a') (q : p ▹ b = b')
: ap (sigma.sigma_functor f g) (sigma_eq p q)
= 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
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)
= sigma_eq (ap f p)
(transport_compose B' f p (g u.1 u.2)⁻¹ ⬝ fn_tr_eq_tr_fn p g u.2⁻¹ ⬝ ap (g v.1) q) :=
: ap (sigma.sigma_functor f g) (sigma_eq p q) =
sigma_eq (ap f p)
((transport_compose B' f p (g u.1 u.2))⁻¹ ⬝ (fn_tr_eq_tr_fn p g u.2)⁻¹ ⬝ ap (g v.1) q) :=
by cases u; cases v; apply ap_sigma_functor_eq_dpair
/- definition 3.11.9(i): Summing up a contractible family of types does nothing. -/
@ -276,7 +268,7 @@ namespace sigma
definition sigma_equiv_of_is_contr_pr1 (B : A → Type) [H : is_contr A] : (Σa, B a) ≃ B (center A)
:=
equiv.mk _ (adjointify
(λu, contr u.1⁻¹ ▹ u.2)
(λu, (contr u.1)⁻¹ ▹ u.2)
(λb, ⟨!center, b⟩)
(λb, ap (λx, x ▹ b) !hprop_eq)
(λu, sigma_eq !contr !tr_inv_tr))

View file

@ -2,6 +2,7 @@
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Module: init.datatypes
Authors: Leonardo de Moura
Basic datatypes

View file

@ -2,7 +2,7 @@
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Module: general_notation
Module: init.reserved_notation
Authors: Leonardo de Moura, Jeremy Avigad
-/
prelude

View file

@ -326,6 +326,7 @@ order for the change to take effect."
("-1p" . ("⁻¹ᵖ"))
("-1e" . ("⁻¹ᵉ"))
("-1h" . ("⁻¹ʰ"))
("-1g" . ("⁻¹ᵍ"))
("qed" . (""))
("x" . ("×"))
("o" . (""))