diff --git a/hott/hit/sphere.hlean b/hott/hit/sphere.hlean index 3322a2bb1..879088f0e 100644 --- a/hott/hit/sphere.hlean +++ b/hott/hit/sphere.hlean @@ -6,9 +6,9 @@ Authors: Floris van Doorn Declaration of the n-spheres -/ -import .suspension +import .suspension types.trunc -open eq nat suspension bool is_trunc unit +open eq nat suspension bool is_trunc unit pointed /- We can define spheres with the following possible indices: @@ -25,6 +25,12 @@ inductive sphere_index : Type₀ := | minus_one : sphere_index | succ : sphere_index → sphere_index +namespace trunc_index + definition sub_one [reducible] (n : sphere_index) : trunc_index := + sphere_index.rec_on n -2 (λ n k, k.+1) + postfix `.-1`:(max+1) := sub_one +end trunc_index + namespace sphere_index /- notation for sphere_index is -1, 0, 1, ... @@ -52,10 +58,19 @@ namespace sphere_index definition empty_of_succ_le_minus_two {n : sphere_index} (H : n .+1 ≤ -1) : empty := H definition of_nat [coercion] [reducible] (n : nat) : sphere_index := - nat.rec_on n (-1.+1) (λ n k, k.+1) + (nat.rec_on n -1 (λ n k, k.+1)).+1 definition trunc_index_of_sphere_index [coercion] [reducible] (n : sphere_index) : trunc_index := - sphere_index.rec_on n -1 (λ n k, k.+1) + (sphere_index.rec_on n -2 (λ n k, k.+1)).+1 + + definition sub_one [reducible] (n : ℕ) : sphere_index := + nat.rec_on n -1 (λ n k, k.+1) + + postfix `.-1`:(max+1) := sub_one + + open trunc_index + definition sub_two_eq_sub_one_sub_one (n : ℕ) : n.-2 = n.-1.-1 := + nat.rec_on n idp (λn p, ap trunc_index.succ p) end sphere_index @@ -66,21 +81,72 @@ definition sphere : sphere_index → Type₀ | n.+1 := suspension (sphere n) namespace sphere + + definition base {n : ℕ} : sphere n := !north + definition pointed_sphere [instance] [constructor] (n : ℕ) : pointed (sphere n) := + pointed.mk base + definition Sphere [constructor] (n : ℕ) : Pointed := pointed.mk' (sphere n) + namespace ops abbreviation S := sphere + notation `S.` := Sphere end ops + open sphere.ops - definition bool_of_sphere [reducible] : sphere 0 → bool := - suspension.rec tt ff (λx, empty.elim x) + definition bool_of_sphere [reducible] : S 0 → bool := + suspension.rec ff tt (λx, empty.elim x) - definition sphere_of_bool [reducible] : bool → sphere 0 - | tt := !north - | ff := !south + definition sphere_of_bool [reducible] : bool → S 0 + | ff := !north + | tt := !south - definition sphere_equiv_bool : sphere 0 ≃ bool := + definition sphere_equiv_bool : S 0 ≃ bool := equiv.MK bool_of_sphere sphere_of_bool (λb, match b with | tt := idp | ff := idp end) (λx, suspension.rec_on x idp idp (empty.rec _)) + definition sphere_eq_bool : S 0 = bool := + ua sphere_equiv_bool + + definition sphere_eq_bool_pointed : S. 0 = Bool := + Pointed_eq sphere_equiv_bool idp + + definition pointed_map_sphere (A : Pointed) (n : ℕ) : map₊ (S. n) A ≃ Ω[n] A := + begin + revert A, induction n with n IH, + { intro A, rewrite [sphere_eq_bool_pointed], apply pointed_map_bool_equiv}, + { intro A, transitivity _, apply suspension_adjoint_loop (S. n) A, apply IH} + end + end sphere + +open sphere sphere.ops + +structure weakly_constant [class] {A B : Type} (f : A → B) := --move + (is_constant : Πa a', f a = f a') + +namespace trunc + open trunc_index + definition is_trunc_of_pointed_map_sphere_constant (n : ℕ) (A : Type) + (H : Π(a : A) (f : map₊ (S. n) (pointed.Mk a)) (x : S n), f x = f base) : is_trunc (n.-2.+1) A := + begin + apply iff.elim_right !is_trunc_iff_is_contr_loop, + intro a, + apply is_trunc_equiv_closed, apply pointed_map_sphere, + fapply is_contr.mk, + { exact pointed_map.mk (λx, a) idp}, + { intro f, fapply pointed_map_eq, + { intro x, esimp, refine !respect_pt⁻¹ ⬝ (!H ⬝ !H⁻¹)}, + { rewrite [▸*,con.right_inv,▸*,con.left_inv]}} + end + + + definition is_trunc_iff_map_sphere_constant (n : ℕ) (A : Type) + (H : Π(f : S n → A) (x : S n), f x = f base) : is_trunc (n.-2.+1) A := + begin + apply is_trunc_of_pointed_map_sphere_constant, + intros, cases f with f p, esimp at *, apply H + end + +end trunc diff --git a/hott/hit/suspension.hlean b/hott/hit/suspension.hlean index 2bb3e783f..3dc240210 100644 --- a/hott/hit/suspension.hlean +++ b/hott/hit/suspension.hlean @@ -6,9 +6,9 @@ Authors: Floris van Doorn Declaration of suspension -/ -import .pushout +import .pushout types.pointed -open pushout unit eq equiv equiv.ops +open pushout unit eq equiv equiv.ops pointed definition suspension (A : Type) : Type := pushout (λ(a : A), star.{0}) (λ(a : A), star.{0}) @@ -76,3 +76,36 @@ attribute suspension.rec suspension.elim [unfold-c 6] [recursor 6] attribute suspension.elim_type [unfold-c 5] attribute suspension.rec_on suspension.elim_on [unfold-c 3] attribute suspension.elim_type_on [unfold-c 2] + +namespace suspension + + definition pointed_suspension [instance] [constructor] (A : Type) : pointed (suspension A) := + pointed.mk !north + + definition suspension_adjoint_loop (A B : Pointed) + : map₊ (pointed.mk' (suspension A)) B ≃ map₊ A (Ω B) := + begin + fapply equiv.MK, + { intro f, fapply pointed_map.mk, + intro a, refine !respect_pt⁻¹ ⬝ ap f (merid a ⬝ (merid pt)⁻¹) ⬝ !respect_pt, + refine ap _ !con.right_inv ⬝ !con.left_inv}, + { intro g, fapply pointed_map.mk, + { esimp, intro a, induction a, + exact pt, + exact pt, + exact g a} , + { reflexivity}}, + { intro g, fapply pointed_map_eq, + intro a, esimp [respect_pt], refine !idp_con ⬝ !ap_con ⬝ ap _ !ap_inv + ⬝ ap _ !elim_merid ⬝ ap _ !elim_merid ⬝ ap _ (respect_pt g) ⬝ _, exact idp, + -- rewrite [ap_con,ap_inv,+elim_merid,idp_con,respect_pt], + esimp, exact sorry}, + { intro f, fapply pointed_map_eq, + { esimp, intro a, induction a; all_goals esimp, + exact !respect_pt⁻¹, + exact !respect_pt⁻¹ ⬝ ap f (merid pt), + apply pathover_eq, exact sorry}, + { esimp, exact !con.left_inv⁻¹}}, + end + +end suspension diff --git a/hott/init/equiv.hlean b/hott/init/equiv.hlean index ad9a3df2b..32da11ddf 100644 --- a/hott/init/equiv.hlean +++ b/hott/init/equiv.hlean @@ -41,7 +41,7 @@ namespace is_equiv protected abbreviation mk [constructor] := @is_equiv.mk' A B f -- The identity function is an equivalence. - definition is_equiv_id (A : Type) : (@is_equiv A A id) := + definition is_equiv_id (A : Type) : (is_equiv (id : A → A)) := is_equiv.mk id id (λa, idp) (λa, idp) (λa, idp) -- The composition of two equivalences is, again, an equivalence. @@ -266,13 +266,15 @@ namespace equiv -- notation for inverse which is not overloaded abbreviation erfl [constructor] := @equiv.refl - definition equiv_of_eq_fn_of_equiv (f : A ≃ B) {f' : A → B} (Heq : f = f') : A ≃ B := + definition equiv_of_eq_fn_of_equiv [constructor] (f : A ≃ B) {f' : A → B} (Heq : f = f') : A ≃ B := equiv.mk f' (is_equiv_eq_closed Heq) - definition eq_equiv_fn_eq (f : A → B) [H : is_equiv f] (a b : A) : (a = b) ≃ (f a = f b) := + --rename: eq_equiv_fn_eq_of_is_equiv + definition eq_equiv_fn_eq [constructor] (f : A → B) [H : is_equiv f] (a b : A) : (a = b) ≃ (f a = f b) := equiv.mk (ap f) !is_equiv_ap - definition eq_equiv_fn_eq_of_equiv (f : A ≃ B) (a b : A) : (a = b) ≃ (f a = f b) := + --rename: eq_equiv_fn_eq + definition eq_equiv_fn_eq_of_equiv [constructor] (f : A ≃ B) (a b : A) : (a = b) ≃ (f a = f b) := equiv.mk (ap f) !is_equiv_ap definition equiv_ap (P : A → Type) {a b : A} (p : a = b) : (P a) ≃ (P b) := diff --git a/hott/init/funext.hlean b/hott/init/funext.hlean index 821c5d8e1..12288a6e7 100644 --- a/hott/init/funext.hlean +++ b/hott/init/funext.hlean @@ -246,6 +246,12 @@ equiv.mk apd10 _ definition eq_of_homotopy [reducible] : f ∼ g → f = g := (@apd10 A P f g)⁻¹ +definition apd10_eq_of_homotopy (p : f ∼ g) : apd10 (eq_of_homotopy p) = p := +right_inv apd10 p + +definition eq_of_homotopy_apd10 (p : f = g) : eq_of_homotopy (apd10 p) = p := +left_inv apd10 p + definition eq_of_homotopy_idp (f : Π x, P x) : eq_of_homotopy (λx : A, idpath (f x)) = idpath f := is_equiv.left_inv apd10 idp diff --git a/hott/init/logic.hlean b/hott/init/logic.hlean index 46250aa5f..c8c65a014 100644 --- a/hott/init/logic.hlean +++ b/hott/init/logic.hlean @@ -181,6 +181,13 @@ namespace iff open eq.ops definition of_eq {a b : Type} (H : a = b) : a ↔ b := iff.intro (λ Ha, H ▸ Ha) (λ Hb, H⁻¹ ▸ Hb) + + definition pi_iff_pi {A : Type} {P Q : A → Type} (H : Πa, (P a ↔ Q a)) : (Πa, P a) ↔ Πa, Q a := + iff.intro (λp a, iff.elim_left (H a) (p a)) (λq a, iff.elim_right (H a) (q a)) + + theorem imp_iff {P : Type} (Q : Type) (p : P) : (P → Q) ↔ Q := + iff.intro (λf, f p) (λq p, q) + end iff attribute iff.refl [refl] @@ -287,7 +294,7 @@ definition decidable_pred [reducible] {A : Type} (R : A → Type) := Π (a definition decidable_rel [reducible] {A : Type} (R : A → A → Type) := Π (a b : A), decidable (R a b) definition decidable_eq [reducible] (A : Type) := decidable_rel (@eq A) definition decidable_ne [instance] {A : Type} [H : decidable_eq A] : decidable_rel (@ne A) := -show ∀ x y : A, decidable (x = y → empty), from _ +show Π x y : A, decidable (x = y → empty), from _ definition ite (c : Type) [H : decidable c] {A : Type} (t e : A) : A := decidable.rec_on H (λ Hc, t) (λ Hnc, e) diff --git a/hott/init/pathover.hlean b/hott/init/pathover.hlean index 8ea90fe6d..c413085f5 100644 --- a/hott/init/pathover.hlean +++ b/hott/init/pathover.hlean @@ -79,6 +79,9 @@ namespace eq definition apdo [unfold-c 6] (f : Πa, B a) (p : a = a₂) : f a =[p] f a₂ := eq.rec_on p idpo + definition oap [unfold-c 6] {C : A → Type} (f : Πa, B a → C a) (p : a = a₂) : f a =[p] f a₂ := + eq.rec_on p idpo + -- infix `⬝` := concato infix `⬝o`:75 := concato -- postfix `⁻¹` := inverseo diff --git a/hott/init/trunc.hlean b/hott/init/trunc.hlean index 63dc90354..31da9581c 100644 --- a/hott/init/trunc.hlean +++ b/hott/init/trunc.hlean @@ -38,8 +38,8 @@ namespace is_trunc definition leq (n m : trunc_index) : Type₀ := trunc_index.rec_on n (λm, unit) (λ n p m, trunc_index.rec_on m (λ p, empty) (λ m q p, p m) p) m - notation x <= y := trunc_index.leq x y - notation x ≤ y := trunc_index.leq x y + infix <= := trunc_index.leq + infix ≤ := trunc_index.leq end trunc_index infix `+2+`:65 := trunc_index.add @@ -51,7 +51,7 @@ namespace is_trunc definition empty_of_succ_le_minus_two {n : trunc_index} (H : n .+1 ≤ -2) : empty := H end trunc_index definition trunc_index.of_nat [coercion] [reducible] (n : nat) : trunc_index := - nat.rec_on n (-1.+1) (λ n k, k.+1) + (nat.rec_on n -2 (λ n k, k.+1)).+2 definition sub_two [reducible] (n : nat) : trunc_index := nat.rec_on n -2 (λ n k, k.+1) diff --git a/hott/types/bool.hlean b/hott/types/bool.hlean index aca68997b..7915d96eb 100644 --- a/hott/types/bool.hlean +++ b/hott/types/bool.hlean @@ -6,7 +6,7 @@ Authors: Floris van Doorn Theorems about the booleans -/ -open is_equiv eq equiv function is_trunc +open is_equiv eq equiv function is_trunc option unit namespace bool @@ -35,4 +35,13 @@ namespace bool definition not_is_hset_type : ¬is_hset Type₀ := assume H : is_hset Type₀, absurd !is_hset.elim eq_bnot_ne_idp + + definition bool_equiv_option_unit : bool ≃ option unit := + begin + fapply equiv.MK, + { intro b, cases b, exact none, exact some star}, + { intro u, cases u, exact ff, exact tt}, + { intro u, cases u with u, reflexivity, cases u, reflexivity}, + { intro b, cases b, reflexivity, reflexivity}, + end end bool diff --git a/hott/types/eq.hlean b/hott/types/eq.hlean index afbc2e29f..e859fcfac 100644 --- a/hott/types/eq.hlean +++ b/hott/types/eq.hlean @@ -7,6 +7,8 @@ Partially ported from Coq HoTT Theorems about path types (identity types) -/ +import .square + open eq sigma sigma.ops equiv is_equiv equiv.ops -- TODO: Rename transport_eq_... and pathover_eq_... to eq_transport_... and eq_pathover_... @@ -19,7 +21,7 @@ namespace eq /- The path spaces of a path space are not, of course, determined; they are just the higher-dimensional structure of the original space. -/ - /- some lemmas about whiskering -/ + /- some lemmas about whiskering or other higher paths -/ definition whisker_left_con_right (p : a1 = a2) {q q' q'' : a2 = a3} (r : q = q') (s : q' = q'') : whisker_left p (r ⬝ s) = whisker_left p r ⬝ whisker_left p s := @@ -51,6 +53,15 @@ namespace eq cases p, cases r, cases q, exact idp end + definition con_right_inv2 (p : a1 = a2) : (con.right_inv p)⁻¹ ⬝ con.right_inv p = idp := + by cases p;exact idp + + definition con_left_inv2 (p : a1 = a2) : (con.left_inv p)⁻¹ ⬝ con.left_inv p = idp := + by cases p;exact idp + + definition ap_eq_ap10 {f g : A → B} (p : f = g) (a : A) : ap (λh, h a) p = ap10 p a := + by cases p;exact idp + /- Transporting in path spaces. There are potentially a lot of these lemmas, so we adopt a uniform naming scheme: @@ -101,6 +112,10 @@ namespace eq -- In the comment we give the fibration of the pathover + definition pathover_eq {f g : A → B} {p : a1 = a2} {q : f a1 = g a1} {r : f a2 = g a2} + (s : square q r (ap f p) (ap g p)) : q =[p] r := + by cases p;apply pathover_idp_of_eq;exact eq_of_vdeg_square s + definition pathover_eq_l (p : a1 = a2) (q : a1 = a3) : q =[p] p⁻¹ ⬝ q := /-(λx, x = a3)-/ by cases p; cases q; exact idpo @@ -160,7 +175,7 @@ namespace eq definition eq_equiv_eq_symm (a1 a2 : A) : (a1 = a2) ≃ (a2 = a1) := equiv.mk inverse _ - definition is_equiv_concat_left [instance] (p : a1 = a2) (a3 : A) + definition is_equiv_concat_left [constructor] [instance] (p : a1 = a2) (a3 : A) : is_equiv (concat p : a2 = a3 → a1 = a3) := is_equiv.mk (concat p) (concat p⁻¹) (con_inv_cancel_left p) @@ -168,10 +183,10 @@ namespace eq (λq, by cases p;cases q;exact idp) local attribute is_equiv_concat_left [instance] - definition equiv_eq_closed_left (a3 : A) (p : a1 = a2) : (a1 = a3) ≃ (a2 = a3) := + definition equiv_eq_closed_left [constructor] (a3 : A) (p : a1 = a2) : (a1 = a3) ≃ (a2 = a3) := equiv.mk (concat p⁻¹) _ - definition is_equiv_concat_right [instance] (p : a2 = a3) (a1 : A) + definition is_equiv_concat_right [constructor] [instance] (p : a2 = a3) (a1 : A) : is_equiv (λq : a1 = a2, q ⬝ p) := is_equiv.mk (λq, q ⬝ p) (λq, q ⬝ p⁻¹) (λq, inv_con_cancel_right q p) @@ -179,10 +194,10 @@ namespace eq (λq, by cases p;cases q;exact idp) local attribute is_equiv_concat_right [instance] - definition equiv_eq_closed_right (a1 : A) (p : a2 = a3) : (a1 = a2) ≃ (a1 = a3) := + definition equiv_eq_closed_right [constructor] (a1 : A) (p : a2 = a3) : (a1 = a2) ≃ (a1 = a3) := equiv.mk (λq, q ⬝ p) _ - definition eq_equiv_eq_closed (p : a1 = a2) (q : a3 = a4) : (a1 = a3) ≃ (a2 = a4) := + definition eq_equiv_eq_closed [constructor] (p : a1 = a2) (q : a3 = a4) : (a1 = a3) ≃ (a2 = a4) := equiv.trans (equiv_eq_closed_left a3 p) (equiv_eq_closed_right a2 q) definition is_equiv_whisker_left (p : a1 = a2) (q r : a2 = a3) diff --git a/hott/types/pathover.hlean b/hott/types/pathover.hlean index dfb291262..833ac67e6 100644 --- a/hott/types/pathover.hlean +++ b/hott/types/pathover.hlean @@ -33,7 +33,7 @@ namespace eq definition pathover_pathover_fl {a' a₂' : A'} {p : a' = a₂'} {a₂ : A} {f : A' → A} {b : Πa, B (f a)} {b₂ : B a₂} {q : Π(a' : A'), f a' = a₂} (r : pathover B (b a') (q a') b₂) (r₂ : pathover B (b a₂') (q a₂') b₂) - (s : squareover B (naturality q p) r r₂ (pathover_ap B f (apdo b p)) (!ap_constant⁻¹ ▸ idpo)) + (s : squareover B (natural_square q p) r r₂ (pathover_ap B f (apdo b p)) (!ap_constant⁻¹ ▸ idpo)) : r =[p] r₂ := by cases p; esimp at s; apply pathover_idp_of_eq; apply eq_of_vdeg_squareover; exact s diff --git a/hott/types/pointed.hlean b/hott/types/pointed.hlean index 214e92fda..727265cce 100644 --- a/hott/types/pointed.hlean +++ b/hott/types/pointed.hlean @@ -6,71 +6,164 @@ Authors: Jakob von Raumer, Floris van Doorn Ported from Coq HoTT -/ -open core is_trunc +import arity .eq .bool .unit .sigma +open is_trunc eq prod sigma nat equiv option is_equiv bool unit structure pointed [class] (A : Type) := (point : A) -namespace pointed - variables {A B : Type} - definition pt [H : pointed A] := point A - - -- Any contractible type is pointed - definition pointed_of_is_contr [instance] (A : Type) [H : is_contr A] : pointed A := - pointed.mk !center - - -- A pi type with a pointed target is pointed - definition pointed_pi [instance] (P : A → Type) [H : Πx, pointed (P x)] - : pointed (Πx, P x) := - pointed.mk (λx, pt) - - -- A sigma type of pointed components is pointed - definition pointed_sigma [instance] (P : A → Type) [G : pointed A] - [H : pointed (P pt)] : pointed (Σx, P x) := - pointed.mk ⟨pt,pt⟩ - - definition pointed_prod [instance] (A B : Type) [H1 : pointed A] [H2 : pointed B] - : pointed (A × B) := - pointed.mk (pt,pt) - - definition pointed_loop [instance] (a : A) : pointed (a = a) := - pointed.mk idp - - definition pointed_fun_closed (f : A → B) [H : pointed A] : pointed B := - pointed.mk (f pt) - -end pointed - structure Pointed := {carrier : Type} (Point : carrier) -open pointed Pointed - -namespace Pointed - attribute carrier [coercion] - protected definition mk' (A : Type) [H : pointed A] : Pointed := Pointed.mk (point A) - definition pointed_carrier [instance] (A : Pointed) : pointed (carrier A) := - pointed.mk (Point A) - - definition Loop_space [reducible] (A : Pointed) : Pointed := - Pointed.mk' (point A = point A) - - definition Iterated_loop_space (A : Pointed) (n : ℕ) : Pointed := - nat.rec_on n A (λn X, Loop_space X) - - prefix `Ω`:95 := Loop_space - notation `Ω[`:95 n:0 `]`:0 A:95 := Iterated_loop_space A n - -end Pointed +open Pointed namespace pointed - export Pointed (hiding cases_on destruct mk) - abbreviation Cases_on := @Pointed.cases_on - abbreviation Destruct := @Pointed.destruct - abbreviation Mk := @Pointed.mk + attribute Pointed.carrier [coercion] + variables {A B : Type} + + definition pt [unfold-c 2] [H : pointed A] := point A + protected abbreviation Mk [constructor] := @Pointed.mk + protected definition mk' [constructor] (A : Type) [H : pointed A] : Pointed := + Pointed.mk (point A) + definition pointed_carrier [instance] [constructor] (A : Pointed) : pointed A := + pointed.mk (Point A) + + -- Any contractible type is pointed + definition pointed_of_is_contr [instance] [constructor] (A : Type) [H : is_contr A] : pointed A := + pointed.mk !center + + -- A pi type with a pointed target is pointed + definition pointed_pi [instance] [constructor] (P : A → Type) [H : Πx, pointed (P x)] + : pointed (Πx, P x) := + pointed.mk (λx, pt) + + -- A sigma type of pointed components is pointed + definition pointed_sigma [instance] [constructor] (P : A → Type) [G : pointed A] + [H : pointed (P pt)] : pointed (Σx, P x) := + pointed.mk ⟨pt,pt⟩ + + definition pointed_prod [instance] [constructor] (A B : Type) [H1 : pointed A] [H2 : pointed B] + : pointed (A × B) := + pointed.mk (pt,pt) + + definition pointed_loop [instance] [constructor] (a : A) : pointed (a = a) := + pointed.mk idp + + definition pointed_bool [instance] [constructor] : pointed bool := + pointed.mk ff + + definition Bool [constructor] : Pointed := + pointed.mk' bool + + definition pointed_fun_closed [constructor] (f : A → B) [H : pointed A] : pointed B := + pointed.mk (f pt) + + definition Loop_space [reducible] [constructor] (A : Pointed) : Pointed := + pointed.mk' (point A = point A) + + -- definition Iterated_loop_space : Pointed → ℕ → Pointed + -- | Iterated_loop_space A 0 := A + -- | Iterated_loop_space A (n+1) := Iterated_loop_space (Loop_space A) n + + definition Iterated_loop_space (A : Pointed) (n : ℕ) : Pointed := + nat.rec_on n (λA, A) (λn IH A, IH (Loop_space A)) A + + prefix `Ω`:(max+1) := Loop_space + notation `Ω[`:95 n:0 `]`:0 A:95 := Iterated_loop_space A n definition iterated_loop_space (A : Type) [H : pointed A] (n : ℕ) : Type := - carrier (Iterated_loop_space (Pointed.mk' A) n) + Ω[n] (pointed.mk' A) + + open equiv.ops + definition Pointed_eq {A B : Pointed} (f : A ≃ B) (p : f pt = pt) : A = B := + begin + cases A with A a, cases B with B b, esimp at *, + fapply apd011 @Pointed.mk, + { apply ua f}, + { rewrite [cast_ua,p]}, + end + + definition add_point [constructor] (A : Type) : Pointed := + Pointed.mk (none : option A) + postfix `₊`:(max+1) := add_point + -- the inclusion A → A₊ is called "some", the extra point "pt" or "none" ("@none A") +end pointed + +open pointed +structure pointed_map (A B : Pointed) := + (map : A → B) (respect_pt : map (Point A) = Point B) + +open pointed_map + +namespace pointed + + abbreviation respect_pt := @pointed_map.respect_pt + + -- definition transport_to_sigma {A B : Pointed} + -- {P : Π(X : Type) (m : X → A → B), (Π(f : X), m f (Point A) = Point B) → (Π(m : A → B), m (Point A) = Point B → X) → Type} + -- (H : P (Σ(f : A → B), f (Point A) = Point B) pr1 pr2 sigma.mk) : + -- P (pointed_map A B) map respect_pt pointed_map.mk := + -- sorry + + + notation `map₊` := pointed_map + attribute pointed_map.map [coercion] + definition pointed_map_eq {A B : Pointed} {f g : map₊ A B} + (r : Πa, f a = g a) (s : respect_pt f = (r pt) ⬝ respect_pt g) : f = g := + begin + cases f with f p, cases g with g q, + esimp at *, + fapply apo011 pointed_map.mk, + { exact eq_of_homotopy r}, + { apply concato_eq, apply pathover_eq_Fl, apply inv_con_eq_of_eq_con, + rewrite [ap_eq_ap10,↑ap10,apd10_eq_of_homotopy,↑respect_pt at *,s]} + end + + definition pointed_map_equiv_left (A : Type) (B : Pointed) : map₊ A₊ B ≃ (A → B) := + begin + fapply equiv.MK, + { intro f a, cases f with f p, exact f (some a)}, + { intro f, fapply pointed_map.mk, + intro a, cases a, exact pt, exact f a, + reflexivity}, + { intro f, reflexivity}, + { intro f, cases f with f p, esimp, fapply pointed_map_eq, + { intro a, cases a; all_goals (esimp at *), exact p⁻¹}, + { esimp, exact !con.left_inv⁻¹}}, + end + + -- set_option pp.notation false + -- definition pointed_map_equiv_right (A : Pointed) (B : Type) + -- : (Σ(b : B), map₊ A (pointed.Mk b)) ≃ (A → B) := + -- begin + -- fapply equiv.MK, + -- { intro u a, cases u with b f, cases f with f p, esimp at f, exact f a}, + -- { intro f, refine ⟨f pt, _⟩, fapply pointed_map.mk, + -- intro a, esimp, exact f a, + -- reflexivity}, + -- { intro f, reflexivity}, + -- { intro u, cases u with b f, cases f with f p, esimp at *, apply sigma_eq p, + -- esimp, apply sorry + -- } + -- end + + + definition pointed_map_bool_equiv (B : Pointed) : map₊ Bool B ≃ B := + begin + fapply equiv.MK, + { intro f, cases f with f p, exact f tt}, + { intro b, fapply pointed_map.mk, + intro u, cases u, exact pt, exact b, + reflexivity}, + { intro b, reflexivity}, + { intro f, cases f with f p, esimp, fapply pointed_map_eq, + { intro a, cases a; all_goals (esimp at *), exact p⁻¹}, + { esimp, exact !con.left_inv⁻¹}}, + end + -- calc + -- map₊ (Pointed.mk' bool) B ≃ map₊ unit₊ B : sorry + -- ... ≃ (unit → B) : pointed_map_equiv + -- ... ≃ B : unit_imp_equiv end pointed diff --git a/hott/types/square.hlean b/hott/types/square.hlean index a0a9bc547..b1bdbc128 100644 --- a/hott/types/square.hlean +++ b/hott/types/square.hlean @@ -91,7 +91,7 @@ namespace eq { intro s, cases s, apply idp}, end - definition naturality [unfold-c 8] {f g : A → B} (p : f ∼ g) (q : a = a') : + definition natural_square [unfold-c 8] {f g : A → B} (p : f ∼ g) (q : a = a') : square (p a) (p a') (ap f q) (ap g q) := eq.rec_on q vrfl diff --git a/hott/types/trunc.hlean b/hott/types/trunc.hlean index dbf0f4361..3d6a6d7f2 100644 --- a/hott/types/trunc.hlean +++ b/hott/types/trunc.hlean @@ -130,8 +130,7 @@ namespace is_trunc section open decidable --this is proven differently in init.hedberg - definition is_hset_of_decidable_eq (A : Type) - [H : decidable_eq A] : is_hset A := + definition is_hset_of_decidable_eq (A : Type) [H : decidable_eq A] : is_hset A := is_hset_of_double_neg_elim (λa b, by_contradiction) end @@ -148,17 +147,43 @@ namespace is_trunc end definition is_trunc_succ_iff_is_trunc_loop (A : Type) (Hn : -1 ≤ n) : - (Π(a : A), is_trunc n (a = a)) ↔ is_trunc (n.+1) A := - iff.intro (is_trunc_succ_of_is_trunc_loop Hn) _ + is_trunc (n.+1) A ↔ Π(a : A), is_trunc n (a = a) := + iff.intro _ (is_trunc_succ_of_is_trunc_loop Hn) - definition is_trunc_iff_is_contr_loop' {n : ℕ} - : (Π(a : A), is_contr (Ω[ n ](Pointed.mk a))) ↔ is_trunc (n.-2.+1) A := +-- set_option pp.implicit true + -- set_option pp.coercions true + -- set_option pp.notation false + definition is_trunc_iff_is_contr_loop_succ (n : ℕ) (A : Type) + : is_trunc n A ↔ Π(a : A), is_contr (Ω[succ n](Pointed.mk a)) := begin - induction n with n H, - { esimp [sub_two,Pointed.Iterated_loop_space], apply iff.intro, - intro H, apply is_hprop_of_imp_is_contr, exact H, - intro H a, exact is_contr_of_inhabited_hprop a}, - { exact sorry}, + revert A, induction n with n IH, + { intros, esimp [Iterated_loop_space], apply iff.intro, + { intros H a, apply is_contr.mk idp, apply is_hprop.elim}, + { intro H, apply is_hset_of_axiom_K, intros, apply is_hprop.elim}}, + { intros, transitivity _, apply @is_trunc_succ_iff_is_trunc_loop n, constructor, + apply iff.pi_iff_pi, intros, + transitivity _, apply IH, + assert H : Πp : a = a, Ω(Pointed.mk p) = Ω(Pointed.mk (idpath a)), + { intros, fapply Pointed_eq, + { esimp, transitivity _, + apply eq_equiv_fn_eq_of_equiv (equiv_eq_closed_right _ p⁻¹), + esimp, apply eq_equiv_eq_closed, apply con.right_inv, apply con.right_inv}, + { esimp, apply con_right_inv2}}, + transitivity _, + apply iff.pi_iff_pi, intro p, + rewrite [↑Iterated_loop_space,↓Iterated_loop_space (Loop_space (pointed.Mk p)) n,H], + apply iff.refl, + apply iff.imp_iff, reflexivity} + end + + definition is_trunc_iff_is_contr_loop (n : ℕ) (A : Type) + : is_trunc (n.-2.+1) A ↔ (Π(a : A), is_contr (Ω[n](pointed.Mk a))) := + begin + cases n with n, + { esimp [sub_two,Iterated_loop_space], apply iff.intro, + intro H a, exact is_contr_of_inhabited_hprop a, + intro H, apply is_hprop_of_imp_is_contr, exact H}, + { apply is_trunc_iff_is_contr_loop_succ}, end end is_trunc open is_trunc diff --git a/hott/types/unit.hlean b/hott/types/unit.hlean new file mode 100644 index 000000000..11b14b193 --- /dev/null +++ b/hott/types/unit.hlean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2015 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn + +Theorems about the booleans +-/ + +open equiv option + +namespace unit + + definition unit_equiv_option_empty : unit ≃ option empty := + begin + fapply equiv.MK, + { intro u, exact none}, + { intro e, exact star}, + { intro e, cases e, reflexivity, contradiction}, + { intro u, cases u, reflexivity}, + end + + definition unit_imp_equiv (A : Type) : (unit → A) ≃ A := + begin + fapply equiv.MK, + { intro f, exact f star}, + { intro a u, exact a}, + { intro a, reflexivity}, + { intro f, apply eq_of_homotopy, intro u, cases u, reflexivity}, + end + +end unit diff --git a/library/logic/connectives.lean b/library/logic/connectives.lean index f2e525120..d0f5bab5a 100644 --- a/library/logic/connectives.lean +++ b/library/logic/connectives.lean @@ -254,6 +254,13 @@ iff.intro (assume H, trivial) (assume H, Ha) theorem iff_self (a : Prop) : (a ↔ a) ↔ true := iff_true_of_self !iff.refl +theorem forall_iff_forall {A : Type} {P Q : A → Prop} (H : ∀a, (P a ↔ Q a)) : (∀a, P a) ↔ ∀a, Q a := +iff.intro (λp a, iff.elim_left (H a) (p a)) (λq a, iff.elim_right (H a) (q a)) + +theorem imp_iff {P : Prop} (Q : Prop) (p : P) : (P → Q) ↔ Q := +iff.intro (λf, f p) (λq p, q) + + /- if-then-else -/ section