feat(hott): more computation rules for trunc_index and use nats for Lemma 8.6.2

This commit is contained in:
Floris van Doorn 2016-02-26 17:00:28 -05:00 committed by Leonardo de Moura
parent e5d5ef9d55
commit c6e628da12
5 changed files with 87 additions and 33 deletions

View file

@ -1,18 +1,18 @@
/- /-
Copyright (c) 2015 Ulrik Buchholtz. All rights reserved. Copyright (c) 2015 Ulrik Buchholtz. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Authors: Ulrik Buchholtz Authors: Ulrik Buchholtz, Floris van Doorn
-/ -/
import types.trunc types.eq types.arrow_2 types.fiber .susp import types.trunc types.eq types.arrow_2 types.fiber .susp
open eq is_trunc is_equiv nat equiv trunc function fiber funext open eq is_trunc is_equiv nat equiv trunc function fiber funext pi
namespace homotopy namespace homotopy
definition is_conn [reducible] (n : trunc_index) (A : Type) : Type := definition is_conn [reducible] (n : ℕ₋₂) (A : Type) : Type :=
is_contr (trunc n A) is_contr (trunc n A)
definition is_conn_equiv_closed (n : trunc_index) {A B : Type} definition is_conn_equiv_closed (n : ℕ₋₂) {A B : Type}
: A ≃ B → is_conn n A → is_conn n B := : A ≃ B → is_conn n A → is_conn n B :=
begin begin
intros H C, intros H C,
@ -21,12 +21,12 @@ namespace homotopy
assumption assumption
end end
definition is_conn_map (n : trunc_index) {A B : Type} (f : A → B) : Type := definition is_conn_map (n : ℕ₋₂) {A B : Type} (f : A → B) : Type :=
Πb : B, is_conn n (fiber f b) Πb : B, is_conn n (fiber f b)
namespace is_conn_map namespace is_conn_map
section section
parameters {n : trunc_index} {A B : Type} {h : A → B} parameters {n : ℕ₋₂} {A B : Type} {h : A → B}
(H : is_conn_map n h) (P : B → n -Type) (H : is_conn_map n h) (P : B → n -Type)
private definition rec.helper : (Πa : A, P (h a)) → Πb : B, trunc n (fiber h b) → P b := private definition rec.helper : (Πa : A, P (h a)) → Πb : B, trunc n (fiber h b) → P b :=
@ -57,14 +57,14 @@ namespace homotopy
end end
section section
parameters {n k : trunc_index} {A B : Type} {f : A → B} parameters {n k : ℕ₋₂} {A B : Type} {f : A → B}
(H : is_conn_map n f) (P : B → (n +2+ k)-Type) (H : is_conn_map n f) (P : B → (n +2+ k)-Type)
include H include H
-- Lemma 8.6.1 -- Lemma 8.6.1
proposition elim_general (t : Πa : A, P (f a)) proposition elim_general : is_trunc_fun k (pi_functor_left f P) :=
: is_trunc k (fiber (λs : (Πb : B, P b), (λa, s (f a))) t) :=
begin begin
intro t,
induction k with k IH, induction k with k IH,
{ apply is_contr_fiber_of_is_equiv, apply is_conn_map.rec, exact H }, { apply is_contr_fiber_of_is_equiv, apply is_conn_map.rec, exact H },
{ apply is_trunc_succ_intro, { apply is_trunc_succ_intro,
@ -98,11 +98,12 @@ namespace homotopy
(@is_trunc_eq (P b) (n +2+ k) (trunctype.struct (P b)) (@is_trunc_eq (P b) (n +2+ k) (trunctype.struct (P b))
(g b) (h b))) } (g b) (h b))) }
end end
end end
section section
universe variables u v universe variables u v
parameters {n : trunc_index} {A : Type.{u}} {B : Type.{v}} {h : A → B} parameters {n : ℕ₋₂} {A : Type.{u}} {B : Type.{v}} {h : A → B}
parameter sec : ΠP : B → trunctype.{max u v} n, parameter sec : ΠP : B → trunctype.{max u v} n,
is_retraction (λs : (Πb : B, P b), λ a, s (h a)) is_retraction (λs : (Πb : B, P b), λ a, s (h a))
@ -127,7 +128,7 @@ namespace homotopy
-- Connectedness is related to maps to and from the unit type, first to -- Connectedness is related to maps to and from the unit type, first to
section section
parameters (n : trunc_index) (A : Type) parameters (n : ℕ₋₂) (A : Type)
definition is_conn_of_map_to_unit definition is_conn_of_map_to_unit
: is_conn_map n (λx : A, unit.star) → is_conn n A := : is_conn_map n (λx : A, unit.star) → is_conn n A :=
@ -161,7 +162,7 @@ namespace homotopy
namespace is_conn namespace is_conn
open pointed unit open pointed unit
section section
parameters {n : trunc_index} {A : Type*} parameters {n : ℕ₋₂} {A : Type*}
[H : is_conn n .+1 A] (P : A → n-Type) [H : is_conn n .+1 A] (P : A → n-Type)
include H include H
@ -180,7 +181,7 @@ namespace homotopy
end end
section section
parameters {n k : trunc_index} {A : Type*} parameters {n k : ℕ₋₂} {A : Type*}
[H : is_conn n .+1 A] (P : A → (n +2+ k)-Type) [H : is_conn n .+1 A] (P : A → (n +2+ k)-Type)
include H include H
@ -223,7 +224,7 @@ namespace homotopy
-- Lemma 7.5.4 -- Lemma 7.5.4
definition retract_of_conn_is_conn [instance] (r : arrow_hom f g) [H : is_retraction r] definition retract_of_conn_is_conn [instance] (r : arrow_hom f g) [H : is_retraction r]
(n : trunc_index) [K : is_conn_map n f] : is_conn_map n g := (n : ℕ₋₂) [K : is_conn_map n f] : is_conn_map n g :=
begin begin
intro b, unfold is_conn, intro b, unfold is_conn,
apply is_contr_retract (trunc_functor n (retraction_on_fiber r b)), apply is_contr_retract (trunc_functor n (retraction_on_fiber r b)),
@ -233,7 +234,7 @@ namespace homotopy
end end
-- Corollary 7.5.5 -- Corollary 7.5.5
definition is_conn_homotopy (n : trunc_index) {A B : Type} {f g : A → B} definition is_conn_homotopy (n : ℕ₋₂) {A B : Type} {f g : A → B}
(p : f ~ g) (H : is_conn_map n f) : is_conn_map n g := (p : f ~ g) (H : is_conn_map n f) : is_conn_map n g :=
@retract_of_conn_is_conn _ _ (arrow.arrow_hom_of_homotopy p) (arrow.is_retraction_arrow_hom_of_homotopy p) n H @retract_of_conn_is_conn _ _ (arrow.arrow_hom_of_homotopy p) (arrow.is_retraction_arrow_hom_of_homotopy p) n H
@ -244,7 +245,7 @@ namespace homotopy
-- Theorem 8.2.1 -- Theorem 8.2.1
open susp open susp
definition is_conn_susp [instance] (n : trunc_index) (A : Type) definition is_conn_susp [instance] (n : ℕ₋₂) (A : Type)
[H : is_conn n A] : is_conn (n .+1) (susp A) := [H : is_conn n A] : is_conn (n .+1) (susp A) :=
is_contr.mk (tr north) is_contr.mk (tr north)
begin begin

View file

@ -7,7 +7,7 @@ The Wedge Sum of Two pType Types
-/ -/
import hit.pointed_pushout .connectedness import hit.pointed_pushout .connectedness
open eq pushout pointed pType unit open eq pushout pointed unit is_trunc.trunc_index pType
definition pwedge (A B : Type*) : Type* := ppushout (pconst punit A) (pconst punit B) definition pwedge (A B : Type*) : Type* := ppushout (pconst punit A) (pconst punit B)
@ -34,18 +34,21 @@ open trunc is_trunc function homotopy
namespace wedge_extension namespace wedge_extension
section section
-- The wedge connectivity lemma (Lemma 8.6.2) -- The wedge connectivity lemma (Lemma 8.6.2)
parameters {A B : Type*} (n m : trunc_index) parameters {A B : Type*} (n m : )
[cA : is_conn n .+2 A] [cB : is_conn m .+2 B] [cA : is_conn n A] [cB : is_conn m B]
(P : A → B → (m .+1 +2+ n .+1)-Type) (P : A → B → (m + n)-Type)
(f : Πa : A, P a (Point B)) (f : Πa : A, P a pt)
(g : Πb : B, P (Point A) b) (g : Πb : B, P pt b)
(p : f (Point A) = g (Point B)) (p : f pt = g pt)
include cA cB include cA cB
private definition Q (a : A) : (n .+1)-Type := private definition Q (a : A) : (n.-1)-Type :=
trunctype.mk trunctype.mk
(fiber (λs : (Πb : B, P a b), s (Point B)) (f a)) (fiber (λs : (Πb : B, P a b), s (Point B)) (f a))
(is_conn.elim_general (P a) (f a)) abstract begin
refine @is_conn.elim_general (m.-1) _ _ _ (λb, trunctype.mk (P a b) _) (f a),
rewrite [-succ_add_succ, of_nat_add_of_nat], intro b, apply trunctype.struct
end end
private definition Q_sec : Πa : A, Q a := private definition Q_sec : Πa : A, Q a :=
is_conn.elim Q (fiber.mk g p⁻¹) is_conn.elim Q (fiber.mk g p⁻¹)

View file

@ -50,7 +50,7 @@ namespace is_trunc
(trunc_index.cases_on n -2 (λn', (trunc_index.cases_on n' -2 id))) (trunc_index.cases_on n -2 (λn', (trunc_index.cases_on n' -2 id)))
(λm', trunc_index.cases_on m' (λm', trunc_index.cases_on m'
(trunc_index.cases_on n -2 id) (trunc_index.cases_on n -2 id)
(trunc_index.rec n (λn' r, succ r))) (add_plus_two n))
/- we give a weird name to the reflexivity step to avoid overloading le.refl /- we give a weird name to the reflexivity step to avoid overloading le.refl
(which can be used if types.trunc is imported) -/ (which can be used if types.trunc is imported) -/
@ -75,6 +75,7 @@ namespace is_trunc
trunc_index.rec_on n nat.zero (λ n k, nat.succ k) trunc_index.rec_on n nat.zero (λ n k, nat.succ k)
postfix `.-2`:(max+1) := sub_two postfix `.-2`:(max+1) := sub_two
postfix `.-1`:(max+1) := λn, (n .-2 .+1)
definition trunc_index.of_nat [coercion] [reducible] (n : ) : ℕ₋₂ := definition trunc_index.of_nat [coercion] [reducible] (n : ) : ℕ₋₂ :=
n.-2.+2 n.-2.+2

View file

@ -174,7 +174,15 @@ namespace pi
/- The functoriality of [forall] is slightly subtle: it is contravariant in the domain type and covariant in the codomain, but the codomain is dependent on the domain. -/ /- The functoriality of [forall] is slightly subtle: it is contravariant in the domain type and covariant in the codomain, but the codomain is dependent on the domain. -/
definition pi_functor [unfold_full] : (Π(a:A), B a) → (Π(a':A'), B' a') := λg a', f1 a' (g (f0 a')) definition pi_functor [unfold_full] : (Π(a:A), B a) → (Π(a':A'), B' a') :=
λg a', f1 a' (g (f0 a'))
definition pi_functor_left [unfold_full] (B : A → Type) : (Π(a:A), B a) → (Π(a':A'), B (f0 a')) :=
pi_functor f0 (λa, id)
definition pi_functor_right [unfold_full] {B' : A → Type} (f1 : Π(a:A), B a → B' a)
: (Π(a:A), B a) → (Π(a:A), B' a) :=
pi_functor id f1
definition ap_pi_functor {g g' : Π(a:A), B a} (h : g ~ g') definition ap_pi_functor {g g' : Π(a:A), B a} (h : g ~ g')
: ap (pi_functor f0 f1) (eq_of_homotopy h) : ap (pi_functor f0 f1) (eq_of_homotopy h)

View file

@ -101,14 +101,55 @@ namespace is_trunc
begin begin
induction n with n IH, induction n with n IH,
{ reflexivity}, { reflexivity},
{ apply ap succ IH} { exact ap succ IH}
end end
definition add_two_sub_two (n : ) : add_two (sub_two n) = n := definition add_two_sub_two (n : ) : add_two (sub_two n) = n :=
begin begin
induction n with n IH, induction n with n IH,
{ reflexivity}, { reflexivity},
{ apply ap nat.succ IH} { exact ap nat.succ IH}
end
definition of_nat_add_plus_two_of_nat (n m : ) : n +2+ m = of_nat (n + m + 2) :=
begin
induction m with m IH,
{ reflexivity},
{ exact ap succ IH}
end
definition of_nat_add_of_nat (n m : ) : of_nat n + of_nat m = of_nat (n + m) :=
begin
induction m with m IH,
{ reflexivity},
{ exact ap succ IH}
end
definition succ_add_plus_two (n m : ℕ₋₂) : n.+1 +2+ m = (n +2+ m).+1 :=
begin
induction m with m IH,
{ reflexivity},
{ exact ap succ IH}
end
definition add_plus_two_succ (n m : ℕ₋₂) : n +2+ m.+1 = (n +2+ m).+1 :=
idp
definition add_succ_succ (n m : ℕ₋₂) : n + m.+2 = n +2+ m :=
idp
definition succ_add_succ (n m : ℕ₋₂) : n.+1 + m.+1 = n +2+ m :=
begin
cases m with m IH,
{ reflexivity},
{ apply succ_add_plus_two}
end
definition succ_succ_add (n m : ℕ₋₂) : n.+2 + m = n +2+ m :=
begin
cases m with m IH,
{ reflexivity},
{ exact !succ_add_succ ⬝ !succ_add_plus_two}
end end
definition succ_sub_two (n : ) : (nat.succ n).-2 = n.-2 .+1 := rfl definition succ_sub_two (n : ) : (nat.succ n).-2 = n.-2 .+1 := rfl