refactor(library/algebra/category/morphism): restore previous (and more readable) proofs

This commit is contained in:
Leonardo de Moura 2015-03-01 06:53:38 -08:00
parent c772d7bf84
commit 25df44ea43

View file

@ -57,19 +57,24 @@ namespace morphism
theorem left_inverse_eq_right_inverse {f : a ⟶ b} {g g' : hom b a} theorem left_inverse_eq_right_inverse {f : a ⟶ b} {g g' : hom b a}
(Hl : g ∘ f = id) (Hr : f ∘ g' = id) : g = g' := (Hl : g ∘ f = id) (Hr : f ∘ g' = id) : g = g' :=
by rewrite [-(id_right g), -Hr, assoc, Hl, id_left] calc
g = g ∘ id : (id_right g)⁻¹
... = g ∘ f ∘ g' : by rewrite -Hr
... = (g ∘ f) ∘ g' : by rewrite assoc
... = id ∘ g' : by rewrite Hl
... = g' : by rewrite id_left
theorem retraction_eq_intro [H : is_section f] (H2 : f ∘ h = id) : retraction_of f = h theorem retraction_eq_intro [H : is_section f] (H2 : f ∘ h = id) : retraction_of f = h :=
:= left_inverse_eq_right_inverse !retraction_compose H2 left_inverse_eq_right_inverse !retraction_compose H2
theorem section_eq_intro [H : is_retraction f] (H2 : h ∘ f = id) : section_of f = h theorem section_eq_intro [H : is_retraction f] (H2 : h ∘ f = id) : section_of f = h :=
:= symm (left_inverse_eq_right_inverse H2 !compose_section) symm (left_inverse_eq_right_inverse H2 !compose_section)
theorem inverse_eq_intro_right [H : is_iso f] (H2 : f ∘ h = id) : f⁻¹ = h theorem inverse_eq_intro_right [H : is_iso f] (H2 : f ∘ h = id) : f⁻¹ = h :=
:= left_inverse_eq_right_inverse !inverse_compose H2 left_inverse_eq_right_inverse !inverse_compose H2
theorem inverse_eq_intro_left [H : is_iso f] (H2 : h ∘ f = id) : f⁻¹ = h theorem inverse_eq_intro_left [H : is_iso f] (H2 : h ∘ f = id) : f⁻¹ = h :=
:= symm (left_inverse_eq_right_inverse H2 !compose_inverse) symm (left_inverse_eq_right_inverse H2 !compose_inverse)
theorem section_eq_retraction (f : a ⟶ b) [Hl : is_section f] [Hr : is_retraction f] : theorem section_eq_retraction (f : a ⟶ b) [Hl : is_section f] [Hr : is_retraction f] :
retraction_of f = section_of f := retraction_of f = section_of f :=
@ -97,14 +102,24 @@ namespace morphism
theorem composition_is_section [instance] [Hf : is_section f] [Hg : is_section g] theorem composition_is_section [instance] [Hf : is_section f] [Hg : is_section g]
: is_section (g ∘ f) := : is_section (g ∘ f) :=
is_section.mk is_section.mk
(show (retraction_of f ∘ retraction_of g) ∘ g ∘ f = id, (calc
by rewrite [-assoc, assoc _ g f, retraction_compose, id_left, retraction_compose]) (retraction_of f ∘ retraction_of g) ∘ g ∘ f
= retraction_of f ∘ retraction_of g ∘ g ∘ f : by rewrite -assoc
... = retraction_of f ∘ (retraction_of g ∘ g) ∘ f : by rewrite (assoc _ g f)
... = retraction_of f ∘ id ∘ f : by rewrite retraction_compose
... = retraction_of f ∘ f : by rewrite id_left
... = id : by rewrite retraction_compose)
theorem composition_is_retraction [instance] [Hf : is_retraction f] [Hg : is_retraction g] theorem composition_is_retraction [instance] [Hf : is_retraction f] [Hg : is_retraction g]
: is_retraction (g ∘ f) := : is_retraction (g ∘ f) :=
is_retraction.mk is_retraction.mk
(show (g ∘ f) ∘ section_of f ∘ section_of g = id, (calc
by rewrite [-assoc, {f ∘ _}assoc, compose_section, id_left, compose_section]) (g ∘ f) ∘ section_of f ∘ section_of g
= g ∘ f ∘ section_of f ∘ section_of g : by rewrite -assoc
... = g ∘ (f ∘ section_of f) ∘ section_of g : by rewrite -assoc
... = g ∘ id ∘ section_of g : by rewrite compose_section
... = g ∘ section_of g : by rewrite id_left
... = id : by rewrite compose_section)
theorem composition_is_inverse [instance] [Hf : is_iso f] [Hg : is_iso g] : is_iso (g ∘ f) := theorem composition_is_inverse [instance] [Hf : is_iso f] [Hg : is_iso g] : is_iso (g ∘ f) :=
!section_retraction_imp_iso !section_retraction_imp_iso
@ -143,8 +158,7 @@ namespace morphism
theorem section_is_mono [instance] (f : a ⟶ b) [H : is_section f] : is_mono f := theorem section_is_mono [instance] (f : a ⟶ b) [H : is_section f] : is_mono f :=
is_mono.mk is_mono.mk
(λ c g h H, (λ c g h H, calc
calc
g = id ∘ g : by rewrite id_left g = id ∘ g : by rewrite id_left
... = (retraction_of f ∘ f) ∘ g : by rewrite -retraction_compose ... = (retraction_of f ∘ f) ∘ g : by rewrite -retraction_compose
... = (retraction_of f ∘ f) ∘ h : by rewrite [-assoc, H, -assoc] ... = (retraction_of f ∘ f) ∘ h : by rewrite [-assoc, H, -assoc]
@ -153,8 +167,7 @@ namespace morphism
theorem retraction_is_epi [instance] (f : a ⟶ b) [H : is_retraction f] : is_epi f := theorem retraction_is_epi [instance] (f : a ⟶ b) [H : is_retraction f] : is_epi f :=
is_epi.mk is_epi.mk
(λ c g h H, (λ c g h H, calc
calc
g = g ∘ id : by rewrite id_right g = g ∘ id : by rewrite id_right
... = g ∘ f ∘ section_of f : by rewrite -compose_section ... = g ∘ f ∘ section_of f : by rewrite -compose_section
... = h ∘ f ∘ section_of f : by rewrite [assoc, H, -assoc] ... = h ∘ f ∘ section_of f : by rewrite [assoc, H, -assoc]
@ -190,23 +203,36 @@ namespace morphism
namespace iso namespace iso
section section
variables {ob : Type} [C : category ob] include C variables {ob : Type} [C : category ob] include C
variables {a b c d : ob} (f : b ⟶ a) variables {a b c d : ob}
(r : c ⟶ d) (q : b ⟶ c) (p : a ⟶ b) variables (f : b ⟶ a) (r : c ⟶ d) (q : b ⟶ c) (p : a ⟶ b)
(g : d ⟶ c) variables (g : d ⟶ c)
variable [Hq : is_iso q] include Hq variable [Hq : is_iso q] include Hq
theorem compose_pV : q ∘ q⁻¹ = id := !compose_inverse theorem compose_pV : q ∘ q⁻¹ = id := !compose_inverse
theorem compose_Vp : q⁻¹ ∘ q = id := !inverse_compose theorem compose_Vp : q⁻¹ ∘ q = id := !inverse_compose
theorem compose_V_pp : q⁻¹ ∘ (q ∘ p) = p := theorem compose_V_pp : q⁻¹ ∘ (q ∘ p) = p :=
by rewrite [assoc, inverse_compose, id_left] calc
q⁻¹ ∘ (q ∘ p) = (q⁻¹ ∘ q) ∘ p : by rewrite assoc
... = id ∘ p : by rewrite inverse_compose
... = p : by rewrite id_left
theorem compose_p_Vp : q ∘ (q⁻¹ ∘ g) = g := theorem compose_p_Vp : q ∘ (q⁻¹ ∘ g) = g :=
by rewrite [assoc, compose_inverse, id_left] calc
q ∘ (q⁻¹ ∘ g) = (q ∘ q⁻¹) ∘ g : by rewrite assoc
... = id ∘ g : by rewrite compose_inverse
... = g : by rewrite id_left
theorem compose_pp_V : (r ∘ q) ∘ q⁻¹ = r := theorem compose_pp_V : (r ∘ q) ∘ q⁻¹ = r :=
by rewrite [-assoc, compose_inverse, id_right] calc
(r ∘ q) ∘ q⁻¹ = r ∘ q ∘ q⁻¹ : by rewrite assoc
... = r ∘ id : by rewrite compose_inverse
... = r : by rewrite id_right
theorem compose_pV_p : (f ∘ q⁻¹) ∘ q = f := theorem compose_pV_p : (f ∘ q⁻¹) ∘ q = f :=
by rewrite [-assoc, inverse_compose, id_right] calc
(f ∘ q⁻¹) ∘ q = f ∘ q⁻¹ ∘ q : by rewrite assoc
... = f ∘ id : by rewrite inverse_compose
... = f : by rewrite id_right
theorem inv_pp [H' : is_iso p] : (q ∘ p)⁻¹ = p⁻¹ ∘ q⁻¹ := theorem inv_pp [H' : is_iso p] : (q ∘ p)⁻¹ = p⁻¹ ∘ q⁻¹ :=
inverse_eq_intro_left inverse_eq_intro_left
@ -220,7 +246,7 @@ namespace morphism
section section
variables {ob : Type} {C : category ob} include C variables {ob : Type} {C : category ob} include C
variables {d c b a : ob} variables {d c b a : ob}
{i : b ⟶ c} {f : b ⟶ a} variables {i : b ⟶ c} {f : b ⟶ a}
{r : c ⟶ d} {q : b ⟶ c} {p : a ⟶ b} {r : c ⟶ d} {q : b ⟶ c} {p : a ⟶ b}
{g : d ⟶ c} {h : c ⟶ b} {g : d ⟶ c} {h : c ⟶ b}
{x : b ⟶ d} {z : a ⟶ c} {x : b ⟶ d} {z : a ⟶ c}