diff --git a/hott/hit/circle.hlean b/hott/hit/circle.hlean index e88d88a7c..dae328191 100644 --- a/hott/hit/circle.hlean +++ b/hott/hit/circle.hlean @@ -6,9 +6,11 @@ Authors: Floris van Doorn Declaration of the circle -/ -import .sphere types.bool types.eq types.int.hott types.arrow types.equiv algebra.fundamental_group algebra.hott +import .sphere +import types.bool types.int.hott types.equiv +import algebra.fundamental_group algebra.hott -open eq suspension bool sphere_index is_equiv equiv equiv.ops is_trunc +open eq suspension bool sphere_index is_equiv equiv equiv.ops is_trunc pi definition circle : Type₀ := sphere 1 @@ -171,12 +173,8 @@ namespace circle begin induction x, { exact loop}, - { exact sorry} + { apply concato_eq, apply pathover_eq_lr, rewrite [con.left_inv,idp_con]} end - -- circle.rec_on x loop - -- (calc - -- loop ▸ loop = loop⁻¹ ⬝ loop ⬝ loop : transport_eq_lr - -- ... = loop : by rewrite [con.left_inv, idp_con]) definition nonidp_neq_idp : nonidp ≠ (λx, idp) := assume H : nonidp = λx, idp, @@ -202,10 +200,8 @@ namespace circle begin induction x, { exact power loop}, - { apply sorry} --- apply eq_of_homotopy, intro a, --- refine !arrow.arrow_transport ⬝ !transport_eq_r ⬝ _, --- rewrite [transport_code_loop_inv,power_con,succ_pred]} + { apply arrow_pathover_left, intro b, apply concato_eq, apply pathover_eq_r, + rewrite [power_con,transport_code_loop]}, end --remove this theorem after #484 @@ -222,9 +218,9 @@ namespace circle apply transport (λ(y : base = base), transport circle.code y _ = _), { exact !power_con_inv ⬝ ap (power loop) !neg_succ⁻¹}, rewrite [▸*,@con_tr _ circle.code,transport_code_loop_inv, ↑[circle.encode] at p, p, -neg_succ]}}, - --{ apply eq_of_homotopy, intro a, apply @is_hset.elim, esimp [circle.code,base,base1], exact _} - { apply sorry} - --simplify after #587 + { apply pathover_of_tr_eq, apply eq_of_homotopy, intro a, apply @is_hset.elim, + esimp [circle.code,base,base1], exact _} + --simplify after #587 end definition circle_eq_equiv (x : circle) : (base = x) ≃ circle.code x := diff --git a/hott/init/pathover.hlean b/hott/init/pathover.hlean index c8f409b60..3865e8aab 100644 --- a/hott/init/pathover.hlean +++ b/hott/init/pathover.hlean @@ -13,7 +13,7 @@ open equiv is_equiv equiv.ops variables {A A' : Type} {B : A → Type} {C : Πa, B a → Type} {a a₂ a₃ a₄ : A} {p : a = a₂} {p₂ : a₂ = a₃} {p₃ : a₃ = a₄} - {b : B a} {b₂ : B a₂} {b₃ : B a₃} {b₄ : B a₄} + {b b' : B a} {b₂ b₂' : B a₂} {b₃ : B a₃} {b₄ : B a₄} {c : C a b} {c₂ : C a₂ b₂} namespace eq @@ -67,6 +67,12 @@ namespace eq definition concato (r : b =[p] b₂) (r₂ : b₂ =[p₂] b₃) : b =[p ⬝ p₂] b₃ := pathover.rec_on r₂ (pathover.rec_on r idpo) + definition concato_eq (r : b =[p] b₂) (q : b₂ = b₂') : b =[p] b₂' := + eq.rec_on q r + + definition eq_concato (q : b = b') (r : b' =[p] b₂) : b =[p] b₂ := + eq.rec_on q (λr, r) r + definition inverseo (r : b =[p] b₂) : b₂ =[p⁻¹] b := pathover.rec_on r idpo @@ -134,6 +140,14 @@ namespace eq from eq.rec_on (eq_of_pathover_idp r) H, left_inv !pathover_idp r ▸ H2 + definition rec_on_right [recursor] {P : Π⦃b₂ : B a₂⦄, b =[p] b₂ → Type} + {b₂ : B a₂} (r : b =[p] b₂) (H : P !pathover_tr) : P r := + by cases r; exact H + + definition rec_on_left [recursor] {P : Π⦃b : B a⦄, b =[p] b₂ → Type} + {b : B a} (r : b =[p] b₂) (H : P !tr_pathover) : P r := + by cases r; exact H + --pathover with fibration B' ∘ f definition pathover_compose (B' : A' → Type) (f : A → A') (p : a = a₂) (b : B' (f a)) (b₂ : B' (f a₂)) : b =[p] b₂ ≃ b =[ap f p] b₂ := diff --git a/hott/types/arrow.hlean b/hott/types/arrow.hlean index 076c899c8..caa579f67 100644 --- a/hott/types/arrow.hlean +++ b/hott/types/arrow.hlean @@ -11,7 +11,7 @@ import types.pi open eq equiv is_equiv funext pi equiv.ops -namespace arrow +namespace pi variables {A A' : Type} {B B' : Type} {C : A → B → Type} {a a' a'' : A} {b b' b'' : B} {f g : A → B} @@ -50,6 +50,34 @@ namespace arrow : (transport (λa, B a → C a) p f) ∼ (λb, p ▸ f (p⁻¹ ▸ b)) := eq.rec_on p (λx, idp) + /- Pathovers -/ + definition arrow_pathover {B C : A → Type} {f : B a → C a} {g : B a' → C a'} {p : a = a'} + (r : Π(b : B a) (b' : B a') (q : b =[p] b'), f b =[p] g b') : f =[p] g := + begin + cases p, apply pathover_idp_of_eq, + apply eq_of_homotopy, intro b, + exact eq_of_pathover_idp (r b b idpo), + end -end arrow + definition arrow_pathover_left {B C : A → Type} {f : B a → C a} {g : B a' → C a'} {p : a = a'} + (r : Π(b : B a), f b =[p] g (p ▸ b)) : f =[p] g := + begin + cases p, apply pathover_idp_of_eq, + apply eq_of_homotopy, intro b, + exact eq_of_pathover_idp (r b), + end + + definition arrow_pathover_right {B C : A → Type} {f : B a → C a} {g : B a' → C a'} {p : a = a'} + (r : Π(b' : B a'), f (p⁻¹ ▸ b') =[p] g b') : f =[p] g := + begin + cases p, apply pathover_idp_of_eq, + apply eq_of_homotopy, intro b, + exact eq_of_pathover_idp (r b), + end + + definition arrow_pathover_constant {B : Type} {C : A → Type} {f : B → C a} {g : B → C a'} + {p : a = a'} (r : Π(b : B), f b =[p] g b) : f =[p] g := + pi_pathover_constant r + +end pi diff --git a/hott/types/eq.hlean b/hott/types/eq.hlean index abc3b240a..afbc2e29f 100644 --- a/hott/types/eq.hlean +++ b/hott/types/eq.hlean @@ -9,6 +9,8 @@ Theorems about path types (identity types) open eq sigma sigma.ops equiv is_equiv equiv.ops +-- TODO: Rename transport_eq_... and pathover_eq_... to eq_transport_... and eq_pathover_... + namespace eq /- Path spaces -/ @@ -99,15 +101,6 @@ namespace eq -- In the comment we give the fibration of the pathover - definition pathover_eq_r_idp (p : a1 = a2) : idp =[p] p := /-(λx, a1 = x)-/ - by cases p; exact idpo - - definition pathover_eq_l_idp (p : a1 = a2) : idp =[p] p⁻¹ := /-(λx, x = a1)-/ - by cases p; exact idpo - - definition pathover_eq_l_idp' (p : a1 = a2) : idp =[p⁻¹] p := /-(λx, x = a2)-/ - by cases p; exact idpo - definition pathover_eq_l (p : a1 = a2) (q : a1 = a3) : q =[p] p⁻¹ ⬝ q := /-(λx, x = a3)-/ by cases p; cases q; exact idpo @@ -139,6 +132,15 @@ namespace eq /-(λx, x = h (f x))-/ by cases p; rewrite [▸*,idp_con];exact idpo + definition pathover_eq_r_idp (p : a1 = a2) : idp =[p] p := /-(λx, a1 = x)-/ + by cases p; exact idpo + + definition pathover_eq_l_idp (p : a1 = a2) : idp =[p] p⁻¹ := /-(λx, x = a1)-/ + by cases p; exact idpo + + definition pathover_eq_l_idp' (p : a1 = a2) : idp =[p⁻¹] p := /-(λx, x = a2)-/ + by cases p; exact idpo + -- The Functorial action of paths is [ap]. /- Equivalences between path spaces -/ diff --git a/hott/types/equiv.hlean b/hott/types/equiv.hlean index 9dc5d1f55..a0acd0d45 100644 --- a/hott/types/equiv.hlean +++ b/hott/types/equiv.hlean @@ -9,7 +9,7 @@ Theorems about the types equiv and is_equiv import .fiber .arrow arity .hprop_trunc -open eq is_trunc sigma sigma.ops arrow pi fiber function equiv +open eq is_trunc sigma sigma.ops pi fiber function equiv namespace is_equiv variables {A B : Type} (f : A → B) [H : is_equiv f] diff --git a/hott/types/pi.hlean b/hott/types/pi.hlean index 825c7b292..62d6c7abf 100644 --- a/hott/types/pi.hlean +++ b/hott/types/pi.hlean @@ -7,7 +7,7 @@ Partially ported from Coq HoTT Theorems about pi-types (dependent function spaces) -/ -import types.sigma +import types.sigma arity open eq equiv is_equiv funext sigma @@ -69,15 +69,30 @@ namespace pi apply eq_of_pathover_idp, apply r end - definition pi_pathover' {C : (Σa, B a) → Type} {f : Πb, C ⟨a, b⟩} {g : Πb', C ⟨a', b'⟩} - {p : a = a'} (r : Π(b : B a) (b' : B a') (q : pathover B b p b'), f b =[dpair_eq_dpair p q] g b') - : f =[p] g := + definition pi_pathover_left {f : Πb, C a b} {g : Πb', C a' b'} {p : a = a'} + (r : Π(b : B a), f b =[apo011 C p !pathover_tr] g (p ▸ b)) : f =[p] g := begin cases p, apply pathover_idp_of_eq, apply eq_of_homotopy, intro b, - apply (@eq_of_pathover_idp _ C), exact (r b b (pathover.idpatho b)), + apply eq_of_pathover_idp, apply r end + definition pi_pathover_right {f : Πb, C a b} {g : Πb', C a' b'} {p : a = a'} + (r : Π(b' : B a'), f (p⁻¹ ▸ b') =[apo011 C p !tr_pathover] g b') : f =[p] g := + begin + cases p, apply pathover_idp_of_eq, + apply eq_of_homotopy, intro b, + apply eq_of_pathover_idp, apply r + end + + definition pi_pathover_constant {C : A → A' → Type} {f : Π(b : A'), C a b} + {g : Π(b : A'), C a' b} {p : a = a'} + (r : Π(b : A'), f b =[p] g b) : f =[p] g := + begin + cases p, apply pathover_idp_of_eq, + apply eq_of_homotopy, intro b, + exact eq_of_pathover_idp (r b), + end /- Maps on paths -/ diff --git a/hott/types/squareover.hlean b/hott/types/squareover.hlean index bba14575f..6121c4940 100644 --- a/hott/types/squareover.hlean +++ b/hott/types/squareover.hlean @@ -6,7 +6,7 @@ Author: Floris van Doorn Theorems about squareovers -/ -import cubical.pathover cubical.square +import types.square open eq equiv is_equiv equiv.ops @@ -41,6 +41,4 @@ namespace cubical definition idsquareo [reducible] [constructor] (b₀₀ : B a₀₀) := @squareover.idsquareo A a₀₀ B b₀₀ definition idso [reducible] [constructor] := @squareover.idsquareo A a₀₀ B b₀₀ - - end cubical