diff --git a/hott/homotopy/connectedness.hlean b/hott/homotopy/connectedness.hlean index 610224375..459c0b0a9 100644 --- a/hott/homotopy/connectedness.hlean +++ b/hott/homotopy/connectedness.hlean @@ -1,18 +1,18 @@ /- Copyright (c) 2015 Ulrik Buchholtz. All rights reserved. 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 -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 - definition is_conn [reducible] (n : trunc_index) (A : Type) : Type := + definition is_conn [reducible] (n : ℕ₋₂) (A : Type) : Type := 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 := begin intros H C, @@ -21,12 +21,12 @@ namespace homotopy assumption 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) namespace is_conn_map 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) 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 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) include H -- Lemma 8.6.1 - proposition elim_general (t : Πa : A, P (f a)) - : is_trunc k (fiber (λs : (Πb : B, P b), (λa, s (f a))) t) := + proposition elim_general : is_trunc_fun k (pi_functor_left f P) := begin + intro t, induction k with k IH, { apply is_contr_fiber_of_is_equiv, apply is_conn_map.rec, exact H }, { apply is_trunc_succ_intro, @@ -98,11 +98,12 @@ namespace homotopy (@is_trunc_eq (P b) (n +2+ k) (trunctype.struct (P b)) (g b) (h b))) } end + end section 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, 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 section - parameters (n : trunc_index) (A : Type) + parameters (n : ℕ₋₂) (A : Type) definition is_conn_of_map_to_unit : is_conn_map n (λx : A, unit.star) → is_conn n A := @@ -161,7 +162,7 @@ namespace homotopy namespace is_conn open pointed unit section - parameters {n : trunc_index} {A : Type*} + parameters {n : ℕ₋₂} {A : Type*} [H : is_conn n .+1 A] (P : A → n-Type) include H @@ -180,7 +181,7 @@ namespace homotopy end 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) include H @@ -223,7 +224,7 @@ namespace homotopy -- Lemma 7.5.4 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 intro b, unfold is_conn, apply is_contr_retract (trunc_functor n (retraction_on_fiber r b)), @@ -233,7 +234,7 @@ namespace homotopy end -- 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 := @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 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) := is_contr.mk (tr north) begin diff --git a/hott/homotopy/wedge.hlean b/hott/homotopy/wedge.hlean index 1c1bd4d7a..0832bbd18 100644 --- a/hott/homotopy/wedge.hlean +++ b/hott/homotopy/wedge.hlean @@ -7,7 +7,7 @@ The Wedge Sum of Two pType Types -/ 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) @@ -34,18 +34,21 @@ open trunc is_trunc function homotopy namespace wedge_extension section -- The wedge connectivity lemma (Lemma 8.6.2) - parameters {A B : Type*} (n m : trunc_index) - [cA : is_conn n .+2 A] [cB : is_conn m .+2 B] - (P : A → B → (m .+1 +2+ n .+1)-Type) - (f : Πa : A, P a (Point B)) - (g : Πb : B, P (Point A) b) - (p : f (Point A) = g (Point B)) + parameters {A B : Type*} (n m : ℕ) + [cA : is_conn n A] [cB : is_conn m B] + (P : A → B → (m + n)-Type) + (f : Πa : A, P a pt) + (g : Πb : B, P pt b) + (p : f pt = g pt) include cA cB - private definition Q (a : A) : (n .+1)-Type := + private definition Q (a : A) : (n.-1)-Type := trunctype.mk (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 := is_conn.elim Q (fiber.mk g p⁻¹) diff --git a/hott/init/trunc.hlean b/hott/init/trunc.hlean index 5aaaf2c46..3f7d83a5d 100644 --- a/hott/init/trunc.hlean +++ b/hott/init/trunc.hlean @@ -29,8 +29,8 @@ namespace is_trunc -/ notation `-1` := trunc_index.succ trunc_index.minus_two -- ISSUE: -1 gets printed as -2.+1? notation `-2` := trunc_index.minus_two - postfix ` .+1`:(max+1) := trunc_index.succ - postfix ` .+2`:(max+1) := λn, (n .+1 .+1) + postfix `.+1`:(max+1) := trunc_index.succ + postfix `.+2`:(max+1) := λn, (n .+1 .+1) notation `ℕ₋₂` := trunc_index -- input using \N-2 definition has_zero_trunc_index [instance] [priority 2000] : has_zero ℕ₋₂ := @@ -50,7 +50,7 @@ namespace is_trunc (trunc_index.cases_on n -2 (λn', (trunc_index.cases_on n' -2 id))) (λm', trunc_index.cases_on m' (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 (which can be used if types.trunc is imported) -/ @@ -64,7 +64,7 @@ namespace is_trunc has_le.mk trunc_index.le attribute trunc_index.add [reducible] - infix `+2+`:65 := trunc_index.add_plus_two + infix ` +2+ `:65 := trunc_index.add_plus_two definition has_add_trunc_index [instance] [priority 2000] : has_add ℕ₋₂ := has_add.mk trunc_index.add @@ -74,7 +74,8 @@ namespace is_trunc definition add_two [reducible] (n : ℕ₋₂) : ℕ := 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 : ℕ) : ℕ₋₂ := n.-2.+2 diff --git a/hott/types/pi.hlean b/hott/types/pi.hlean index 2ff7606a4..28dfa92c9 100644 --- a/hott/types/pi.hlean +++ b/hott/types/pi.hlean @@ -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. -/ - 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') : ap (pi_functor f0 f1) (eq_of_homotopy h) diff --git a/hott/types/trunc.hlean b/hott/types/trunc.hlean index f9b1cc51a..2018b967b 100644 --- a/hott/types/trunc.hlean +++ b/hott/types/trunc.hlean @@ -101,14 +101,55 @@ namespace is_trunc begin induction n with n IH, { reflexivity}, - { apply ap succ IH} + { exact ap succ IH} end definition add_two_sub_two (n : ℕ) : add_two (sub_two n) = n := begin induction n with n IH, { 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 definition succ_sub_two (n : ℕ) : (nat.succ n).-2 = n.-2 .+1 := rfl