diff --git a/_CoqProject b/_CoqProject index 4ae3a0b3e8d..b6569647ffe 100644 --- a/_CoqProject +++ b/_CoqProject @@ -137,6 +137,8 @@ theories/Spaces/Torus/TorusHomotopy.v theories/Modalities/ReflectiveSubuniverse.v theories/Modalities/Modality.v theories/Modalities/Accessible.v +theories/Modalities/Descent.v +theories/Modalities/Separated.v theories/Modalities/Lex.v theories/Modalities/Topological.v theories/Modalities/Notnot.v diff --git a/contrib/HoTTBook.v b/contrib/HoTTBook.v index b26506d20b4..365643ea4c6 100644 --- a/contrib/HoTTBook.v +++ b/contrib/HoTTBook.v @@ -454,7 +454,7 @@ Definition Book_3_6_2 `{Funext} (A : Type) (B : A -> Type) (* ================================================== thm:prop-equiv-trunc *) (** Lemma 3.9.1 *) -Definition Book_3_9_1 := @HoTT.Truncations.Core.TrM.RSU.isequiv_to_O_inO (-1)%trunc. +Definition Book_3_9_1 := @HoTT.Modalities.ReflectiveSubuniverse.isequiv_to_O_inO (Tr (-1)). (* ================================================== cor:UC *) (** Corollary 3.9.2 *) @@ -490,7 +490,7 @@ Definition Book_3_11_6 := @HoTT.Types.Forall.trunc_forall. (** Lemma 3.11.7 *) Definition Book_3_11_7a := @HoTT.Idempotents.contr_retracttype. -Definition Book_3_11_7 := @HoTT.Truncations.Core.TrM.RSU.inO_to_O_retract (-2)%trunc. +Definition Book_3_11_7 := @HoTT.Modalities.ReflectiveSubuniverse.inO_to_O_retract (Tr (-2)). (* ================================================== thm:contr-paths *) (** Lemma 3.11.8 *) @@ -639,7 +639,7 @@ Definition Book_4_6_1 := @HoTT.Basics.Trunc.IsTruncMap (-1). (* ================================================== thm:mono-surj-equiv *) (** Theorem 4.6.3 *) -Definition Book_4_6_3 := @HoTT.Truncations.Core.TrM.RSU.isequiv_conn_ino_map. +Definition Book_4_6_3 := @HoTT.Modalities.ReflectiveSubuniverse.isequiv_conn_ino_map. (* ================================================== thm:two-out-of-three *) (** Theorem 4.7.1 *) @@ -862,7 +862,7 @@ Definition Book_6_9_1 := @HoTT.Truncations.Core.Trunc.Trunc_ind 0. (* ================================================== thm:trunc0-lump *) (** Lemma 6.9.2 *) -Definition Book_6_9_2 := @HoTT.Truncations.Core.TrM.isequiv_o_to_O. +Definition Book_6_9_2 := @HoTT.Modalities.ReflectiveSubuniverse.isequiv_o_to_O. (* ================================================== thm:set-pushout *) (** Lemma 6.9.3 *) @@ -1107,7 +1107,7 @@ Definition Book_7_3_12 := @HoTT.Truncations.Core.equiv_path_Tr. (* ================================================== cor:totrunc-is-connected *) (** Corollary 7.5.8 *) -Definition Book_7_5_8 := @HoTT.Truncations.Core.TrM.conn_map_to_O. +Definition Book_7_5_8 := @HoTT.Modalities.Modality.conn_map_to_O. (* ================================================== thm:nconn-to-ntype-const *) (** Corollary 7.5.9 *) diff --git a/contrib/HoTTBookExercises.v b/contrib/HoTTBookExercises.v index cee1273452d..030f25ce863 100644 --- a/contrib/HoTTBookExercises.v +++ b/contrib/HoTTBookExercises.v @@ -1563,7 +1563,7 @@ End Book_7_1. (* ================================================== ex:double-negation-modality *) (** Exercise 7.12 *) -Definition Book_7_12 := HoTT.Modalities.Notnot.Notnot. +Definition Book_7_12 := @HoTT.Modalities.Notnot.NotNot. (* ================================================== ex:prop-modalities *) (** Exercise 7.13 *) diff --git a/theories/Algebra/AbelianGroup.v b/theories/Algebra/AbelianGroup.v index dc6ff3e258d..d7eb4d453b0 100644 --- a/theories/Algebra/AbelianGroup.v +++ b/theories/Algebra/AbelianGroup.v @@ -4,7 +4,6 @@ Require Import HIT.Coeq. Require Import Algebra.Group. Require Import Algebra.Subgroup. Require Import Cubical. -Import TrM. Local Open Scope mc_mult_scope. diff --git a/theories/Algebra/Aut.v b/theories/Algebra/Aut.v index 39a2a09f089..ab7cc050715 100644 --- a/theories/Algebra/Aut.v +++ b/theories/Algebra/Aut.v @@ -4,7 +4,6 @@ Require Import Types. Require Import Truncations. Require Import Factorization. Require Import Algebra.ooGroup. -Import TrM. Local Open Scope path_scope. @@ -23,13 +22,13 @@ Definition equiv_baut_image_unit X Proof. unfold BAut, image; simpl. apply equiv_functor_sigma_id; intros Z; simpl. - apply equiv_O_functor; unfold hfiber. + apply Trunc_functor_equiv; unfold hfiber. refine ((equiv_contr_sigma _)^-1 oE _). apply equiv_path_inverse. Defined. (** Now we can define [Aut X], by proving that [BAut X] is connected. *) -Definition Aut (X : Type@{i}) : ooGroup@{j u a}. +Definition Aut (X : Type) : ooGroup. Proof. refine (Build_ooGroup (Build_pType { Z : Type & merely (Z = X) } (X ; tr 1)) _). diff --git a/theories/Algebra/ooGroup.v b/theories/Algebra/ooGroup.v index ec145ba63bd..9a3a2dd45f2 100644 --- a/theories/Algebra/ooGroup.v +++ b/theories/Algebra/ooGroup.v @@ -5,8 +5,6 @@ Require Import Pointed. Require Import Truncations. Require Import Colimits.Quotient. -Import TrM. - Local Open Scope path_scope. (** Keyed unification makes [rewrite !loops_functor_group] take a really long time. See https://coq.inria.fr/bugs/show_bug.cgi?id=4544 for more discussion. *) @@ -19,8 +17,8 @@ Local Unset Keyed Unification. (** ** Definition *) Record ooGroup := - { classifying_space : pType@{i} ; - isconn_classifying_space : IsConnected@{u a i} 0 classifying_space + { classifying_space : pType ; + isconn_classifying_space : IsConnected 0 classifying_space }. Global Existing Instance isconn_classifying_space. diff --git a/theories/Classes/interfaces/canonical_names.v b/theories/Classes/interfaces/canonical_names.v index 682ce6d7440..985fb184a4d 100644 --- a/theories/Classes/interfaces/canonical_names.v +++ b/theories/Classes/interfaces/canonical_names.v @@ -398,7 +398,7 @@ Class Bind (M : Type -> Type) := bind : forall {A B}, M A -> (A -> M B) -> M B. Class Enumerable@{i} (A : Type@{i}) := { enumerator : nat -> A ; enumerator_issurj :> - TrM.RSU.IsConnMap@{Uhuge Ularge i i Ularge} (trunc_S minus_two) enumerator }. + IsConnMap@{i} (trunc_S minus_two) enumerator }. Arguments enumerator A {_} _. Arguments enumerator_issurj A {_} _. diff --git a/theories/Classes/theory/rationals.v b/theories/Classes/theory/rationals.v index 46c09b394b3..81def348833 100644 --- a/theories/Classes/theory/rationals.v +++ b/theories/Classes/theory/rationals.v @@ -556,7 +556,7 @@ destruct (le_or_lt (enumerator Q n) 0) as [E|E]. Defined. Lemma Qpos_is_enumerator : - TrM.RSU.IsConnMap@{Uhuge Ularge UQ UQ Ularge} (trunc_S minus_two) Qpos_enumerator. + IsConnMap@{UQ} (trunc_S minus_two) Qpos_enumerator. Proof. apply BuildIsSurjection. unfold hfiber. diff --git a/theories/Comodalities/CoreflectiveSubuniverse.v b/theories/Comodalities/CoreflectiveSubuniverse.v index e93dd110cfd..ec3ca28f948 100644 --- a/theories/Comodalities/CoreflectiveSubuniverse.v +++ b/theories/Comodalities/CoreflectiveSubuniverse.v @@ -2,7 +2,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import UnivalenceImpliesFunext. Require Import Modalities.Modality Modalities.Open. -Import OpM. Local Open Scope path_scope. diff --git a/theories/Constant.v b/theories/Constant.v index 0fa5db47adf..adf3c689b4e 100644 --- a/theories/Constant.v +++ b/theories/Constant.v @@ -3,7 +3,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import HProp FunextVarieties. Require Import Extensions Factorization Modalities.Modality. Require Import HoTT.Truncations. -Import TrM. Local Open Scope path_scope. Local Open Scope trunc_scope. diff --git a/theories/DProp.v b/theories/DProp.v index b559b577e48..bb02480feec 100644 --- a/theories/DProp.v +++ b/theories/DProp.v @@ -5,7 +5,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import TruncType HProp UnivalenceImpliesFunext. Require Import HoTT.Truncations. -Import TrM. Local Open Scope path_scope. diff --git a/theories/Extensions.v b/theories/Extensions.v index 532a5009c61..00982808727 100644 --- a/theories/Extensions.v +++ b/theories/Extensions.v @@ -18,33 +18,29 @@ Section Extensions. (* TODO: consider naming for [ExtensionAlong] and subsequent lemmas. As a name for the type itself, [Extension] or [ExtensionAlong] seems great; but resultant lemma names such as [path_extension] (following existing naming conventions) are rather misleading. *) (** This elimination rule (and others) can be seen as saying that, given a fibration over the codomain and a section of it over the domain, there is some *extension* of this to a section over the whole codomain. It can also be considered as an equivalent form of an [hfiber] of precomposition-with-[f] that replaces paths by pointwise paths, thereby avoiding [Funext]. *) - - Definition ExtensionAlong {A B : Type} (f : A -> B) - (P : B -> Type) (d : forall x:A, P (f x)) - := { s : forall y:B, P y & forall x:A, s (f x) = d x }. - (** [ExtensionAlong] takes 5 universe parameters: + Definition ExtensionAlong@{a b p m} {A : Type@{a}} {B : Type@{b}} (f : A -> B) + (P : B -> Type@{p}) (d : forall x:A, P (f x)) + := (* { s : forall y:B, P y & forall x:A, s (f x) = d x }. *) + sig@{m m} (fun (s : forall y:B, P y) => forall x:A, s (f x) = d x). + (** [ExtensionAlong] takes 4 universe parameters: - the size of A - the size of B - the size of P - - >= max(B,P) - - >= max(A,P). - The following [Check] verifies that this is in fact the case. *) - (** We would like to say [Check], but because of bug #4517, https://coq.inria.fr/bugs/show_bug.cgi?id=4517, we can't. *) - Definition check_ExtensionAlong@{a b p m n} : True@{Set}. - Proof. - Check ExtensionAlong@{a b p m n}. - Abort. - (** If necessary, we could coalesce the latter two with a universe annotation, but that would make the definition harder to read. *) + - >= max(A,B,P) + *) - (** It's occasionally useful to be able to modify those max universes. *) - Definition lift_extensionalong@{a b p m1 n1 m2 n2} {A : Type@{a}} {B : Type@{b}} (f : A -> B) - (P : B -> Type@{p}) (d : forall x:A, P (f x)) - : ExtensionAlong@{a b p m1 n1} f P d -> ExtensionAlong@{a b p m2 n2} f P d - := fun ext => (ext.1 ; ext.2). + (** It's occasionally useful to be able to modify those universes. For each of the universes [a], [b], [p], we give an initial one, a final one, and a "minimum" one smaller than both and where the type actually lives. *) + Definition lift_extensionalong@{a1 a2 amin b1 b2 bmin p1 p2 pmin m1 m2} {A : Type@{amin}} {B : Type@{bmin}} (f : A -> B) + (P : B -> Type@{pmin}) (d : forall x:A, P (f x)) + : ExtensionAlong@{a1 b1 p1 m1} f P d -> ExtensionAlong@{a2 b2 p2 m2} f P d. + Proof. + intros ext. + (** If we just give [ext], it will collapse the universes. (Anyone stepping through this proof should do [Set Printing Universes] and look at the universes to see that they're different in [ext] and in the goal.) So we decompose [ext] into two components and give them separately. *) + assert (e2 := ext.2); set (e1 := ext.1) in e2. + cbn in e2. (** Curiously, without this line we get a spurious universe inequality [p1 <= m2]. *) + exact (e1;e2). + Defined. (** We called it [lift_extensionalong], but in fact it doesn't require the new universes to be bigger than the old ones, only that they both satisfy the max condition. *) - Definition lower_extensionalong@{a b p e0 e1} {A : Type@{a}} {B : Type@{b}} (f : A -> B) - (P : B -> Type@{p}) (d : forall x:A, P (f x)) - := lift_extensionalong@{a b p e0 e1 e0 e1} f P d. Definition equiv_path_extension `{Funext} {A B : Type} {f : A -> B} {P : B -> Type} {d : forall x:A, P (f x)} @@ -97,7 +93,7 @@ Section Extensions. := match n with | 0 => Unit | S n => (forall (g : forall a, C (f a)), - ExtensionAlong@{i j k l l} f C g) * + ExtensionAlong@{i j k l} f C g) * forall (h k : forall b, C b), ExtendableAlong n f (fun b => h b = k b) end. @@ -106,36 +102,24 @@ Section Extensions. - size of B - size of C - size of result (>= A,B,C) *) - (** We would like to say [Check], but because of bug #4517, https://coq.inria.fr/bugs/show_bug.cgi?id=4517, we can't. *) - Definition check_ExtendableAlong@{a b c r} : True@{Set}. - Proof. - Check ExtendableAlong@{a b c r}. - Abort. Global Arguments ExtendableAlong n%nat_scope {A B}%type_scope (f C)%function_scope. (** We can modify the universes, as with [ExtensionAlong]. *) - Definition lift_extendablealong@{i j k l1 l2} - (n : nat) {A : Type@{i}} {B : Type@{j}} - (f : A -> B) (C : B -> Type@{k}) - : ExtendableAlong@{i j k l1} n f C -> ExtendableAlong@{i j k l2} n f C. + Definition lift_extendablealong@{a1 a2 amin b1 b2 bmin p1 p2 pmin m1 m2} + (n : nat) {A : Type@{amin}} {B : Type@{bmin}} + (f : A -> B) (P : B -> Type@{pmin}) + : ExtendableAlong@{a1 b1 p1 m1} n f P -> ExtendableAlong@{a2 b2 p2 m2} n f P. Proof. - revert C; simple_induction n n IH; intros C. + revert P; simple_induction n n IH; intros P. - intros _; exact tt. - intros ext; split. - + intros g; exact (lift_extensionalong _ _ _ (fst ext g)). - + intros h k; exact (IH _ (snd ext h k)). + + intros g; exact (lift_extensionalong@{a1 a2 amin b1 b2 bmin p1 p2 pmin m1 m2} _ _ _ (fst ext g)). + + intros h k. + (** Unles we give the universe explicitly here, [kmin] gets collapsed to [k1]. *) + pose (P' := (fun b => h b = k b) : B -> Type@{pmin}). + exact (IH P' (snd ext h k)). Defined. - (** We would like to say [Check], but because of bug #4517, https://coq.inria.fr/bugs/show_bug.cgi?id=4517, we can't. *) - Definition check_lift_extendablealong@{i j k l1 l2} : True@{Set}. - Proof. - Check lift_extendablealong@{i j k l1 l2}. - Abort. - - Definition lower_extendablealong@{i j k e} - (n : nat) {A : Type@{i}} {B : Type@{j}} - (f : A -> B) (C : B -> Type@{k}) - := lift_extendablealong@{i j k e e} n f C. Definition equiv_extendable_pathsplit `{Funext} (n : nat) {A B : Type} (C : B -> Type) (f : A -> B) @@ -340,29 +324,15 @@ Section Extensions. (f : A -> B) (C : B -> Type@{k}) : Type@{l} := forall n : nat, ExtendableAlong@{i j k l} n f C. (** Universe parameters are the same as for [ExtendableAlong]. *) - (** We would like to say [Check], but because of bug #4517, https://coq.inria.fr/bugs/show_bug.cgi?id=4517, we can't. *) - Definition check_ooExtendableAlong@{a b c r} : True@{Set}. - Proof. - Check ooExtendableAlong@{a b c r}. - Abort. Global Arguments ooExtendableAlong {A B}%type_scope (f C)%function_scope. (** Universe modification. *) - Definition lift_ooextendablealong@{i j k l1 l2} - {A : Type@{i}} {B : Type@{j}} - (f : A -> B) (C : B -> Type@{k}) - : ooExtendableAlong@{i j k l1} f C -> ooExtendableAlong@{i j k l2} f C - := fun ext n => lift_extendablealong n f C (ext n). - (** We would like to say [Check], but because of bug #4517, https://coq.inria.fr/bugs/show_bug.cgi?id=4517, we can't. *) - Definition check_ooextendablealong@{i j k l1 l2} : True@{Set}. - Proof. - Check lift_ooextendablealong@{i j k l1 l2}. - Abort. - Definition lower_ooextendablealong@{i j k e1 e2} - {A : Type@{i}} {B : Type@{j}} - (f : A -> B) (C : B -> Type@{k}) - := lift_ooextendablealong@{i j k e1 e2} f C. + Definition lift_ooextendablealong@{a1 a2 amin b1 b2 bmin p1 p2 pmin m1 m2} + {A : Type@{amin}} {B : Type@{bmin}} + (f : A -> B) (P : B -> Type@{pmin}) + : ooExtendableAlong@{a1 b1 p1 m1} f P -> ooExtendableAlong@{a2 b2 p2 m2} f P + := fun ext n => lift_extendablealong@{a1 a2 amin b1 b2 bmin p1 p2 pmin m1 m2} n f P (ext n). (** We take part of the data from [ps 1] and part from [ps 2] so that the inverse chosen is the expected one. *) Definition isequiv_ooextendable `{Funext} diff --git a/theories/Fibrations.v b/theories/Fibrations.v index 0438a30ef1b..e126a8c3197 100644 --- a/theories/Fibrations.v +++ b/theories/Fibrations.v @@ -307,6 +307,10 @@ Proof. intros x y. apply isequiv_fcontr,_. Defined. +Definition equiv_ap_isembedding {A B} (f : A -> B) `{!IsEmbedding f} (x y : A) + : (x = y) <~> (f x = f y) + := Build_Equiv _ _ (ap f) _. + Definition isembedding_isequiv_ap {A B} (f : A -> B) `{!forall x y, IsEquiv (@ap _ _ f x y)} : IsEmbedding f. Proof. diff --git a/theories/HIT/Circle.v b/theories/HIT/Circle.v index cff0845dcb8..6f99dae5eec 100644 --- a/theories/HIT/Circle.v +++ b/theories/HIT/Circle.v @@ -10,7 +10,6 @@ Require Import HIT.Coeq. Require Import Modalities.Modality Truncations. Require Import Cubical.DPath. -Import TrM. Local Open Scope path_scope. Generalizable Variables X A B f g n. @@ -201,6 +200,7 @@ Proof. refine (trunc_equiv' (n := 0) Int equiv_loopS1_int^-1). Defined. + (** ** Iteration of equivalences *) (** If [P : S1 -> Type] is defined by a type [X] and an autoequivalence [f], then the image of [n:Int] regarded as in [base = base] is [iter_int f n]. *) diff --git a/theories/HIT/FreeIntQuotient.v b/theories/HIT/FreeIntQuotient.v index effe0d2654b..496a5c3f5a3 100644 --- a/theories/HIT/FreeIntQuotient.v +++ b/theories/HIT/FreeIntQuotient.v @@ -3,7 +3,6 @@ Require Import HoTT.Basics HoTT.Types UnivalenceImpliesFunext. Require Import TruncType HProp HSet. Require Import Spaces.Int. Require Import HIT.Coeq HIT.Circle HIT.Flattening Truncations. -Import TrM. Local Open Scope path_scope. diff --git a/theories/HIT/Spheres.v b/theories/HIT/Spheres.v index 32bb671b547..7a90cde34fb 100644 --- a/theories/HIT/Spheres.v +++ b/theories/HIT/Spheres.v @@ -8,7 +8,6 @@ Require Import HProp NullHomotopy. Require Import Homotopy.Suspension HIT.Circle HIT.TwoSphere. Require Import Pointed.pSusp Pointed.Core. Require Import Truncations. -Import TrM. Local Open Scope pointed_scope. Local Open Scope trunc_scope. diff --git a/theories/HIT/surjective_factor.v b/theories/HIT/surjective_factor.v index 1967ca2f1d6..ec17fee2ec5 100644 --- a/theories/HIT/surjective_factor.v +++ b/theories/HIT/surjective_factor.v @@ -21,7 +21,7 @@ Section surjective_factor. Qed. Definition surjective_factor_aux := - @TrM.RSU.conn_map_elim + @conn_map_elim _ _ _ _ Esurj (fun b => exists c : C, forall a, g a = b -> f a = c) ishprop_surjective_factor_aux (fun a => exist (fun c => forall a, _ -> _ = c) (f a) (fun a' => Eg a' a)). diff --git a/theories/HoTT.v b/theories/HoTT.v index c2be72ddf66..062d8597420 100644 --- a/theories/HoTT.v +++ b/theories/HoTT.v @@ -79,12 +79,14 @@ Require Export HoTT.Colimits.Colimit_Sigma. Require Export HoTT.Modalities.ReflectiveSubuniverse. Require Export HoTT.Modalities.Modality. Require Export HoTT.Modalities.Accessible. -Require Export HoTT.Modalities.Lex. -Require Export HoTT.Modalities.Topological. Require Export HoTT.Modalities.Notnot. Require Export HoTT.Modalities.Identity. Require Export HoTT.Modalities.Localization. Require Export HoTT.Modalities.Nullification. +Require Export HoTT.Modalities.Descent. +Require Export HoTT.Modalities.Separated. +Require Export HoTT.Modalities.Lex. +Require Export HoTT.Modalities.Topological. Require Export HoTT.Modalities.Open. Require Export HoTT.Modalities.Closed. Require Export HoTT.Modalities.Fracture. diff --git a/theories/Homotopy/BlakersMassey.v b/theories/Homotopy/BlakersMassey.v index 0dfdb4c2de8..49dc5752e8b 100644 --- a/theories/Homotopy/BlakersMassey.v +++ b/theories/Homotopy/BlakersMassey.v @@ -40,11 +40,7 @@ Defined. (** ** Setup *) -Module GenBlakersMassey (Os : ReflectiveSubuniverses). - Import Os. - Module Import Os_Theory := ReflectiveSubuniverses_Theory Os. - - Section GBM. +Section GBM. Context {X Y : Type} (Q : X -> Y -> Type). (** Here's the hypothesis of ABFJ generalized Blakers-Massey. It works for any reflective subuniverse, not only modalities! *) @@ -600,14 +596,10 @@ Now we claim that the left-hand map of this span is also an equivalence. Rather (** This version is sufficient for the classical Blakers-Massey theorem, as we'll see below, since its leg-wise connectivity hypothesis implies the above surjectivity assumption. ABFJ have a different method for eliminating the surjectivity assumption using a lemma about pushouts of monos also being pullbacks, though it seems to only work for coderight. *) - End GBM. -End GenBlakersMassey. +End GBM. (** ** The classical Blakers-Massey Theorem *) -Import TrM. -Module Import BlakersMassey := GenBlakersMassey Truncation_RSUs. - Global Instance blakers_massey `{Univalence} (m n : trunc_index) {X Y : Type} (Q : X -> Y -> Type) `{forall y, IsConnected m.+1 { x : X & Q x y } } @@ -618,5 +610,4 @@ Proof. intros r. srefine (contr_code_inhab Q (m +2+ n) _ x (merely_isconnected n _) (spushr Q y) r). - intros x1 x3 y2 y4 q12 q32 q34; apply isconnected_join; exact _. Defined. diff --git a/theories/Homotopy/ClassifyingSpace.v b/theories/Homotopy/ClassifyingSpace.v index 80dc516c0a0..9cbfd42ed1d 100644 --- a/theories/Homotopy/ClassifyingSpace.v +++ b/theories/Homotopy/ClassifyingSpace.v @@ -8,7 +8,6 @@ Require Import Homotopy.HSpace. Require Import TruncType. Require Import Truncations. Require Import UnivalenceImpliesFunext. -Import TrM. Local Open Scope pointed_scope. Local Open Scope trunc_scope. diff --git a/theories/Homotopy/EMSpace.v b/theories/Homotopy/EMSpace.v index 892665d6e1a..78d9ff091d2 100644 --- a/theories/Homotopy/EMSpace.v +++ b/theories/Homotopy/EMSpace.v @@ -10,7 +10,6 @@ Require Import Homotopy.ClassifyingSpace. Require Import Homotopy.HSpace. Require Import TruncType. Require Import UnivalenceImpliesFunext. -Import TrM. (* Formalisation of Eilenberg-MacLane spaces *) @@ -244,7 +243,7 @@ Section EilenbergMacLane. apply ap, ap. destruct n; reflexivity. } destruct p. - apply conn_map_to_O. } + rapply conn_map_to_O. } all: reflexivity. Defined. diff --git a/theories/Homotopy/Freudenthal.v b/theories/Homotopy/Freudenthal.v index 3c94cacf18c..16af11f885d 100644 --- a/theories/Homotopy/Freudenthal.v +++ b/theories/Homotopy/Freudenthal.v @@ -7,8 +7,6 @@ Require Import Homotopy.Join. Require Import Homotopy.Suspension. Require Import Homotopy.BlakersMassey. -Import TrM. - (** * The Freudenthal Suspension Theorem *) (** The Freudenthal suspension theorem is a fairly trivial corollary of the Blakers-Massey theorem. The only real work is to relate the span-pushout that we used for Blakers-Massey to the naive pushout that we used to define suspension. *) diff --git a/theories/Homotopy/HSpace.v b/theories/Homotopy/HSpace.v index bfcb374c1db..613d87e97ed 100644 --- a/theories/Homotopy/HSpace.v +++ b/theories/Homotopy/HSpace.v @@ -2,7 +2,6 @@ Require Import Basics. Require Import Types. Require Import Truncations UnivalenceImpliesFunext. Require Export Classes.interfaces.abstract_algebra. -Import TrM. Local Open Scope trunc_scope. Local Open Scope mc_mult_scope. diff --git a/theories/Homotopy/HomotopyGroup.v b/theories/Homotopy/HomotopyGroup.v index ddd3cc9f583..5067924191b 100644 --- a/theories/Homotopy/HomotopyGroup.v +++ b/theories/Homotopy/HomotopyGroup.v @@ -7,8 +7,6 @@ Require Import Truncations. Require Import Spaces.Nat. Require Import Modalities.ReflectiveSubuniverse. -Import TrM. - Local Open Scope pointed_scope. Local Open Scope path_scope. @@ -146,7 +144,7 @@ Proof. - cbn; apply Trunc_functor_compose. - etransitivity. + apply O_functor_homotopy, iterated_loops_functor_compose. - + refine (O_functor_compose 0%trunc _ _ x). + + refine (O_functor_compose (Tr 0) _ _ x). Defined. Definition pi_2functor (n : nat) diff --git a/theories/Homotopy/Join.v b/theories/Homotopy/Join.v index 77ca31f848f..35118c237ec 100644 --- a/theories/Homotopy/Join.v +++ b/theories/Homotopy/Join.v @@ -8,7 +8,6 @@ Require Import NullHomotopy. Require Import Extensions. Require Import Colimits.Pushout. Require Import Truncations. -Import TrM. Local Open Scope path_scope. diff --git a/theories/Homotopy/Pi1S1.v b/theories/Homotopy/Pi1S1.v index 1a4e3f42ced..e7d4398638d 100644 --- a/theories/Homotopy/Pi1S1.v +++ b/theories/Homotopy/Pi1S1.v @@ -10,7 +10,6 @@ Require Import HIT.Spheres. Require Import Truncations. Require Import Algebra.Z. Require Import Truncations. -Import TrM. (* Calculation of Pi 1 S1 *) @@ -48,4 +47,4 @@ Section Pi1S1. apply pequiv_pSph1_to_S1. Defined. -End Pi1S1. \ No newline at end of file +End Pi1S1. diff --git a/theories/Homotopy/Suspension.v b/theories/Homotopy/Suspension.v index 5b8739fc3a7..9931c239dca 100644 --- a/theories/Homotopy/Suspension.v +++ b/theories/Homotopy/Suspension.v @@ -5,9 +5,9 @@ Require Import Cubical. Require Import WildCat. Require Import Colimits.Pushout. Require Import NullHomotopy. -Require Import Truncations. +Require Import Truncations.Core. +Require Import Modalities.ReflectiveSubuniverse Modalities.Modality. Require Import Extensions. -Import TrM. (** * The suspension of a type *) @@ -532,7 +532,7 @@ Proof. apply isconnected_from_elim. intros C H' f. exists (f North). assert ({ p0 : f North = f South & forall x:X, ap f (merid x) = p0 }) - as [p0 allpath_p0] by (apply (isconnected_elim n); apply H'). + as [p0 allpath_p0] by (apply (isconnected_elim n); rapply H'). apply (Susp_ind (fun a => f a = f North) 1 p0^). intros x. apply (concat (transport_paths_Fl _ _)). diff --git a/theories/Homotopy/WhiteheadsPrinciple.v b/theories/Homotopy/WhiteheadsPrinciple.v index bd6f6ada37e..d4466d1b820 100644 --- a/theories/Homotopy/WhiteheadsPrinciple.v +++ b/theories/Homotopy/WhiteheadsPrinciple.v @@ -6,7 +6,6 @@ Require Import EquivalenceVarieties UnivalenceImpliesFunext. Require Import Algebra.Group. Require Import Homotopy.HomotopyGroup. Require Import Truncations. -Import TrM. Local Open Scope pointed_scope. Local Open Scope nat_scope. @@ -85,7 +84,7 @@ Proof. - pose proof (ii x 0) as h2. unfold pi_functor in h2; cbn in h2. refine (@isequiv_homotopic _ _ _ _ h2 _). - apply O_functor_homotopy; intros p. + apply (O_functor_homotopy (Tr 0)); intros p. exact (concat_1p _ @ concat_p1 _). - intros p k; revert p. assert (h3 : forall (y:A) (q:x=y), diff --git a/theories/Idempotents.v b/theories/Idempotents.v index 09d5249ba9b..53d1355882f 100644 --- a/theories/Idempotents.v +++ b/theories/Idempotents.v @@ -3,7 +3,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import Fibrations FunextVarieties UnivalenceImpliesFunext EquivalenceVarieties Constant. Require Import HoTT.Truncations. Require Import PathAny. -Import TrM. Local Open Scope nat_scope. Local Open Scope path_scope. diff --git a/theories/Limits/Pullback.v b/theories/Limits/Pullback.v index df6af12f391..48d979897f1 100644 --- a/theories/Limits/Pullback.v +++ b/theories/Limits/Pullback.v @@ -1,6 +1,6 @@ (* -*- mode: coq; mode: visual-line -*- *) Require Import HoTT.Basics HoTT.Types. -Require Import Fibrations Cubical.PathSquare. +Require Import Fibrations EquivalenceVarieties Cubical.PathSquare. Local Open Scope path_scope. @@ -64,6 +64,16 @@ Definition equiv_ispullback {A B C D} : A <~> Pullback k g := Build_Equiv _ _ (pullback_corec p) ip. +(** This is equivalent to the transposed square being a pullback. *) +Definition ispullback_symm {A B C D} + {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + (p : g o h == k o f) (pb : IsPullback (fun a => (p a)^)) + : IsPullback p. +Proof. + rapply (cancelL_isequiv (equiv_pullback_symm g k)). + apply pb. +Defined. + (** The pullback of the projections [{d:D & P d} -> D <- {d:D & Q d}] is equivalent to [{d:D & P d * Q d}]. *) Definition ispullback_sigprod {D : Type} (P Q : D -> Type) : IsPullback (fun z:{d:D & P d * Q d} => 1%path : (z.1;fst z.2).1 = (z.1;snd z.2).1). @@ -80,6 +90,42 @@ Definition equiv_sigprod_pullback {D : Type} (P Q : D -> Type) : {d:D & P d * Q d} <~> Pullback (@pr1 D P) (@pr1 D Q) := Build_Equiv _ _ _ (ispullback_sigprod P Q). +(** For any commutative square, the fiber of the fibers is equivalent to the fiber of the "gap map" [pullback_corec]. *) +Definition hfiber_pullback_corec {A B C D} + {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + (p : k o f == g o h) (b : B) (c : C) (q : k b = g c) + : hfiber (pullback_corec p) (b; c; q) <~> hfiber (functor_hfiber p b) (c; q^). +Proof. + unfold hfiber, functor_hfiber, functor_sigma. + refine (equiv_sigma_assoc _ _ oE _). + apply equiv_functor_sigma_id; intros a; cbn. + refine (_ oE (equiv_path_sigma _ _ _)^-1); cbn. + apply equiv_functor_sigma_id; intro p0; cbn. + rewrite transport_sigma'; cbn. + refine ((equiv_path_sigma _ _ _) oE _ oE (equiv_path_sigma _ _ _)^-1); cbn. + apply equiv_functor_sigma_id; intro p1; cbn. + rewrite !transport_paths_Fr, !transport_paths_Fl. + refine (_ oE (equiv_ap (equiv_path_inverse _ _) _ _)); cbn. + apply equiv_concat_l. + refine (_ @ (inv_pp _ _)^). apply whiskerL. + refine (_ @ (inv_pp _ _)^). apply whiskerL. + symmetry; apply inv_V. +Defined. + +(** If the induced map on fibers is an equivalence, then a square is a pullback. *) +Definition ispullback_isequiv_functor_hfiber {A B C D} + {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + (p : k o f == g o h) + (e : forall b:B, IsEquiv (functor_hfiber p b)) + : IsPullback p. +Proof. + unfold IsPullback. + apply isequiv_fcontr; intro x. + rapply contr_equiv'. + - symmetry; apply hfiber_pullback_corec. + - apply fcontr_isequiv, e. +Defined. + (** The pullback of a map along another one *) Definition pullback_along {A B C} (f : B -> A) (g : C -> A) : Pullback f g -> B diff --git a/theories/Modalities/Accessible.v b/theories/Modalities/Accessible.v index f776ef8ff78..06f5ab8c0e0 100644 --- a/theories/Modalities/Accessible.v +++ b/theories/Modalities/Accessible.v @@ -16,7 +16,7 @@ Local Open Scope path_scope. We now give the basic definitions related to accessibility, using [ooExtendableAlong] as our notion of equivalence as we did with reflective subuniverses. The actual construction of a reflective subuniverse by localization will be in [Localization]. *) -Record LocalGenerators := +Record LocalGenerators@{a} := { lgen_indices : Type@{a} ; lgen_domain : lgen_indices -> Type@{a} ; lgen_codomain : lgen_indices -> Type@{a} ; @@ -31,57 +31,40 @@ Module Import IsLocal_Internal. (forall (i : lgen_indices f), ooExtendableAlong (f i) (fun _ => X)). End IsLocal_Internal. -Module Type Accessible_ReflectiveSubuniverses - (Os : ReflectiveSubuniverses). - - Export Os. - - (** In examples (such as localization), the reason we need the extra universe parameter [a] is that it describes the size of the generators. Therefore, here we intentionally collapse that parameter with the parameter of [LocalGenerators]. *) - Parameter acc_gen : ReflectiveSubuniverse@{u a} -> LocalGenerators@{a}. - Check acc_gen@{u a}. (** Verify that we have the right number of universes *) - - Parameter inO_iff_islocal - : forall (O : ReflectiveSubuniverse@{u a}) (X : Type@{i}), +Class IsAccRSU@{a i} (O : Subuniverse@{i}) := +{ + acc_lgen : LocalGenerators@{a} ; + inO_iff_islocal : forall (X : Type@{i}), (** We call [iff] explicitly to control the number of universe parameters. *) - iff@{i i i} - (In@{u a i} O X) - (IsLocal@{i i a} (acc_gen@{u a} O) X). - Check inO_iff_islocal@{u a i}. - -End Accessible_ReflectiveSubuniverses. - -Module Accessible_ReflectiveSubuniverses_Theory - (Os : ReflectiveSubuniverses) - (Acc : Accessible_ReflectiveSubuniverses Os). + iff@{i i i} (In O X) (IsLocal acc_lgen X) ; +}. - Import Os Acc. - Module Import Os_Theory := ReflectiveSubuniverses_Theory Os. +Arguments acc_lgen O {_}. +Arguments inO_iff_islocal O {_} X. - Definition O_inverts_generators {O : ReflectiveSubuniverse} - (i : lgen_indices (acc_gen O)) - : O_inverts O (acc_gen O i). + Global Instance O_inverts_generators {O : ReflectiveSubuniverse} `{IsAccRSU O} + (i : lgen_indices (acc_lgen O)) + : O_inverts O (acc_lgen O i). Proof. - pose (ext_dom := fst (inO_iff_islocal O (O (lgen_domain (acc_gen O) i))) _). - pose (ext_cod := fst (inO_iff_islocal O (O (lgen_codomain (acc_gen O) i))) _). + pose (ext_dom := fst (inO_iff_islocal O (O (lgen_domain (acc_lgen O) i))) _). + pose (ext_cod := fst (inO_iff_islocal O (O (lgen_codomain (acc_lgen O) i))) _). simple refine (isequiv_adjointify _ _ _ _). - apply O_rec. exact ((fst (ext_dom i 1%nat) (to O _)).1). - apply O_indpaths; intros x; simpl. rewrite O_rec_beta. refine ((fst (snd (ext_cod i 2) - (fun x => O_functor O (acc_gen O i) + (fun x => O_functor O (acc_lgen O i) ((fst (ext_dom i 1%nat) (to O _)).1 x)) _) _).1 x); intros a. rewrite ((fst (ext_dom i 1%nat) (to O _)).2 a). apply to_O_natural. - apply O_indpaths; intros x; simpl. - rewrite (to_O_natural O (acc_gen O i) x). + rewrite (to_O_natural O (acc_lgen O i) x). rewrite O_rec_beta. apply ((fst (ext_dom i 1%nat) (to O _)).2 x). Qed. -End Accessible_ReflectiveSubuniverses_Theory. - (** The construction of the localization reflective subuniverse for any family of maps will be in [Localization]. *) @@ -131,37 +114,26 @@ Definition ooextendable_isnull_fibers {A B} (f : A -> B) (C : B -> Type) := fun null n => extendable_isnull_fibers n f C null. (** We define a modality to be accessible if it consists of the null types for some family of generators as above. *) -Module Type Accessible_Modalities (Os : Modalities). - Import Os. - Module Import Os_Theory := Modalities_Theory Os. - - (** See comment in [Accessible_ReflectiveSubuniverses] about collapsing universes. *) - Parameter acc_gen : Modality@{u a} -> NullGenerators@{a}. - Check acc_gen@{u a}. (** Verify that we have the right number of universes *) - - Parameter inO_iff_isnull - : forall (O : Modality@{u a}) (X : Type@{i}), - iff@{i i i} - (In@{u a i} O X) - (IsNull@{a i} (acc_gen@{u a} O) X). - Check inO_iff_isnull@{u a i}. -End Accessible_Modalities. +Class IsAccModality@{a i} (O : Subuniverse@{i}) := +{ + acc_ngen : NullGenerators@{a} ; + inO_iff_isnull : forall (X : Type@{i}), + iff@{i i i} (In O X) (IsNull acc_ngen X) ; +}. -Module Accessible_Modalities_Theory - (Os : Modalities) - (Acc : Accessible_Modalities Os). +Arguments acc_ngen O {_}. +Arguments inO_iff_isnull O {_} X. - Export Os Acc. - Module Export Os_Theory := Modalities_Theory Os. - Import Acc.Os_Theory. +Section AccessibleModalities. + Context (O : Modality) {acco : IsAccModality O}. (** Unsurprisingly, the generators are connected. *) - Global Instance isconnected_acc_gen O i : IsConnected O (acc_gen O i). + Global Instance isconnected_acc_ngen i : IsConnected O (acc_ngen O i). Proof. apply isconnected_from_elim_to_O. - pose (H := fst (fst (inO_iff_isnull O (O (acc_gen O i))) _ i 1%nat) - (to O ((acc_gen O) i))). + pose (H := fst (fst (inO_iff_isnull O (O (acc_ngen O i))) _ i 1%nat) + (to O ((acc_ngen O) i))). exists (H.1 tt). exact (fun x => (H.2 x)^). Defined. @@ -169,22 +141,22 @@ Module Accessible_Modalities_Theory (** If all the generators are inhabited, some things become a bit simpler. *) Section InhabitedGenerators. - Context (O : Modality) (inhab : forall i, acc_gen O i). + Context (inhab : forall i, acc_ngen O i). (** For testing modal-ness of types, it suffices for all maps out of a generator to be constant. Probably one could do without [Funext]. *) Definition inO_const_fromgen `{Funext} A - (c : forall i (f : acc_gen O i -> A), NullHomotopy f) + (c : forall i (f : acc_ngen O i -> A), NullHomotopy f) : In O A. Proof. apply (snd (inO_iff_isnull O A)); intros i. - apply (equiv_ooextendable_isequiv _ _)^-1. - simple refine (isequiv_adjointify _ _ _ _). + apply ((equiv_ooextendable_isequiv _ _)^-1%equiv). + snrapply isequiv_adjointify. - intros f []; exact (c i f).1. - intros f; apply path_arrow; intros x. simpl; unfold composeD. symmetry; exact ((c i f).2 x). - intros g; apply path_arrow; intros []. - pose ((c i (g oD (null_to_local_generators (acc_gen O)) i)).2). + pose ((c i (g oD (null_to_local_generators (acc_ngen O)) i)).2). simpl in p; unfold composeD in p. symmetry; apply p, inhab. Defined. @@ -200,139 +172,49 @@ Module Accessible_Modalities_Theory End InhabitedGenerators. -End Accessible_Modalities_Theory. +End AccessibleModalities. -(** We will now show that a modality is accessible in this sense if and only if its underlying reflective subuniverse is accessible in the sense previously defined. These proofs involve a bit of annoying module wrangling. Fortunately, we (almost?) never need to actually use them; in practice accessible modalities usually seem to be given to us with the appropriate sort of generators. *) +(** We will now show that a modality is accessible in this sense if and only if its underlying reflective subuniverse is accessible in the sense previously defined. We (almost?) never need to actually use this, though; in practice accessible modalities usually seem to be given to us with the appropriate sort of generators. *) (** One direction of this implication is trivial. *) -Module Accessible_Modalities_to_ReflectiveSubuniverses - (Os : Modalities) (Acc : Accessible_Modalities Os). - - (** Coq won't let us write [<: Accessible_ReflectiveSubuniverses (Modalities_to_ReflectiveSubuniverses Os)]; it says "Application of modules is restricted to paths" (a "path" being something like [Foo.Bar.Baz]). Thus, every intermediate module has to be given its own name. *) - Module Os_RSU := Modalities_to_ReflectiveSubuniverses Os. - Module AccRSU <: Accessible_ReflectiveSubuniverses Os_RSU. - - Import Os_RSU Acc. +Global Instance acc_rsu_modality (O : Modality) `{IsAccModality O} + : IsAccRSU O + := Build_IsAccRSU O (null_to_local_generators (acc_ngen O)) (fun X => inO_iff_isnull O X). - Definition acc_gen : ReflectiveSubuniverse@{u a} -> LocalGenerators@{a} - := fun (O : ReflectiveSubuniverse@{u a}) => - (null_to_local_generators (acc_gen O)). - - Definition inO_iff_islocal@{u a i} - : forall (O : ReflectiveSubuniverse@{u a}) (X : Type@{i}), - iff@{i i i} - (In@{u a i} O X) - (IsLocal@{i i a} (acc_gen@{u a} O) X) - := inO_iff_isnull@{u a i}. - - End AccRSU. -End Accessible_Modalities_to_ReflectiveSubuniverses. - -(** The converse is less trivial. *) -Module Accessible_Modalities_from_ReflectiveSubuniverses - (Os : Modalities). - - Module Os_RSU := Modalities_to_ReflectiveSubuniverses Os. - Module AccMod (Acc : Accessible_ReflectiveSubuniverses Os_RSU) - <: Accessible_Modalities Os. - - Import Os Acc. - Module Import Os_Theory := Modalities_Theory Os. - Module Import Acc_Theory := Accessible_ReflectiveSubuniverses_Theory Os_RSU Acc. - - (** The idea is as follows. By [ooextendable_isnull_fibers], we can detect locality with respect to a map by nullity with respect to its fibers. Therefore, our first thought might be to just consider all the fibers of all the maps that we are localizing at. However, this doesn't quite work because [ooextendable_isnull_fibers] is not an if-and-only-if, so not every modal type would necessarily be null for that type family. +(** For the less trivial converse, the idea is as follows. By [ooextendable_isnull_fibers], we can detect locality with respect to a map by nullity with respect to its fibers. Therefore, our first thought might be to just consider all the fibers of all the maps that we are localizing at. However, this doesn't quite work because [ooextendable_isnull_fibers] is not an if-and-only-if, so not every modal type would necessarily be null for that type family. We do know, however, that if [f] is an [O]-connected map, then any [O]-modal type is null for its fibers (since they are [O]-connected types). There is no *a priori* why all the maps we localize at should end up being connected for the modality; they will always be inverted, but not every inverted map is connected (unless the modality is lex). But if [f : A -> B] is [O]-inverted, then the [O]-connected map [to O A] is (up to equivalence) the composite of [f] with the [O]-connected map [to O B]. Thus, if [X] is null for the fibers of [to O A] and [to O B], it will be [f]-local and hence [O]-modal, while all [O]-modal types will be null for these fibers since they are connected. *) - Definition acc_gen (O : Modality@{u a}) : NullGenerators@{a}. - Proof. - refine (Build_NullGenerators - ( { i : lgen_indices@{a} (acc_gen O) - & O (lgen_domain@{a} (acc_gen O) i) } - + { i : lgen_indices@{a} (acc_gen O) - & O (lgen_codomain@{a} (acc_gen O) i) }) - _). - intros [ [i x] | [i x] ]; exact (hfiber (to O _) x). - Defined. - - Definition inO_iff_isnull@{u a i} (O : Modality@{u a}) (X : Type@{i}) - : iff@{i i i} (In@{u a i} O X) (IsNull@{a i} (acc_gen O) X). - Proof. - pose proof (@conn_map_to_O@{u a a a}). - split. - - intros X_inO [ [i x] | [i x] ]; - exact (ooextendable_const_isconnected_inO@{u a u a a i i} O _ _). - (* MS: annotation can be dropped *) - - intros Xnull. - apply (snd (inO_iff_islocal O X)); intros i. - refine (cancelL_ooextendable@{a a a i i i i i i} - (fun _ => X) (Acc.acc_gen O i) - (to O (lgen_codomain (Acc.acc_gen O) i)) _ _). - + apply ooextendable_isnull_fibers@{a a i i a a i a a i}; intros x. - exact (Xnull (inr (i;x))). - + refine (ooextendable_homotopic _ - (O_functor O (Acc.acc_gen O i) - o to O (lgen_domain (Acc.acc_gen O) i)) _ _). - 1:apply to_O_natural. - apply ooextendable_compose@{a a a i i i i}. - * apply ooextendable_equiv, O_inverts_generators. - * apply ooextendable_isnull_fibers@{a a i i a a i a a i}; intros x. - exact (Xnull (inl (i;x))). - Defined. - - End AccMod. -End Accessible_Modalities_from_ReflectiveSubuniverses. +(** We don't make this an [Instance] since it is rarely used, and would cause loops when combined with the previous one. *) +Definition acc_modality_rsu (O : Modality) `{IsAccRSU O} + : IsAccModality O. +Proof. + unshelve econstructor. + { refine (Build_NullGenerators + ( { i : lgen_indices@{a} (acc_lgen O) + & O (lgen_domain@{a} (acc_lgen O) i) } + + { i : lgen_indices@{a} (acc_lgen O) + & O (lgen_codomain@{a} (acc_lgen O) i) }) + _). + intros [ [i x] | [i x] ]; exact (hfiber (to O _) x). } + { assert (cm := @conn_map_to_O O). + split. + - intros X_inO [ [i x] | [i x] ]; + exact (ooextendable_const_isconnected_inO O _ _). + - intros Xnull. + apply (snd (inO_iff_islocal O X)); intros i. + refine (cancelL_ooextendable (fun _ => X) (acc_lgen O i) + (to O (lgen_codomain (acc_lgen O) i)) _ _). + + apply ooextendable_isnull_fibers; intros x. + exact (Xnull (inr (i;x))). + + refine (ooextendable_homotopic _ + (O_functor O (acc_lgen O i) + o to O (lgen_domain (acc_lgen O) i)) _ _). + 1:apply to_O_natural. + apply ooextendable_compose. + * apply ooextendable_equiv, O_inverts_generators. + * apply ooextendable_isnull_fibers; intros x. + exact (Xnull (inl (i;x))). } +Defined. (** The construction of the nullification modality for any family of types will be in [Nullification]. *) - -(** ** Restrictions are accessible *) - -Module Accessible_Restriction_ReflectiveSubuniverses - (Os : ReflectiveSubuniverses) - (Acc : Accessible_ReflectiveSubuniverses Os) - (Res : ReflectiveSubuniverses_Restriction_Data Os). - - Module Import New <: ReflectiveSubuniverses - := ReflectiveSubuniverses_Restriction Os Res. - - Module Accessible_New <: Accessible_ReflectiveSubuniverses New. - - Definition acc_gen : New.ReflectiveSubuniverse@{u a} -> LocalGenerators@{a} - := fun O => Acc.acc_gen (Res.ReflectiveSubuniverses_restriction O). - - Definition inO_iff_islocal - : forall (O : New.ReflectiveSubuniverse@{u a}) (X : Type@{i}), - iff@{i i i} - (In@{u a i} O X) - (IsLocal@{i i a} (acc_gen@{u a} O) X) - := fun O => Acc.inO_iff_islocal (Res.ReflectiveSubuniverses_restriction O). - - End Accessible_New. - -End Accessible_Restriction_ReflectiveSubuniverses. - -Module Accessible_Restriction_Modalities - (Os : Modalities) - (Acc : Accessible_Modalities Os) - (Res : Modalities_Restriction_Data Os). - - Module Import New <: Modalities - := Modalities_Restriction Os Res. - - Module Accessible_New <: Accessible_Modalities New. - - Module Import Os_Theory := Modalities_Theory New. - - Definition acc_gen : New.Modality@{u a} -> NullGenerators@{a} - := fun O => Acc.acc_gen (Res.Modalities_restriction O). - - Definition inO_iff_isnull - : forall (O : New.Modality@{u a}) (X : Type@{i}), - iff@{i i i} - (In@{u a i} O X) - (IsNull@{a i} (acc_gen@{u a} O) X) - := fun O => Acc.inO_iff_isnull (Res.Modalities_restriction O). - - End Accessible_New. - -End Accessible_Restriction_Modalities. diff --git a/theories/Modalities/Closed.v b/theories/Modalities/Closed.v index 1f4c7aba5eb..4b83ed6dc40 100644 --- a/theories/Modalities/Closed.v +++ b/theories/Modalities/Closed.v @@ -43,151 +43,63 @@ Section ClosedModalTypes. End ClosedModalTypes. -Record Closed_Modality := Cl { unCl : hProp }. - (** Exercise 7.13(ii): Closed modalities *) -Module ClosedModalities <: Modalities. - - Definition Modality : Type@{u} := Closed_Modality@{a}. - - Definition O_reflector : Modality@{u a} -> Type@{i} -> Type@{i} - := fun O X => Join@{a i i} (unCl O) X. - - Definition In : Modality@{u a} -> Type@{i} -> Type@{i} - := fun O X => unCl O -> Contr X. - - Definition O_inO (O : Modality@{u a}) (T : Type@{i}) - : In@{u a i} O (O_reflector@{u a i} O T). - Proof. - intros u. - pose (contr_inhabited_hprop _ u). - exact _. - Defined. - - Definition to (O : Modality@{u a}) (T : Type@{i}) - : T -> O_reflector@{u a i} O T - := fun x => push (inr x). - - Definition inO_equiv_inO (O : Modality@{u a}) (T : Type@{i}) (U : Type@{j}) - (T_inO : In@{u a i} O T) (f : T -> U) (feq : IsEquiv f) - : let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - In@{u a j} O U. - Proof. +Definition Cl (U : hProp) : Modality. +Proof. + snrapply Build_Modality. + - intros X; exact (U -> Contr X). + - exact _. + - intros T B T_inO f feq. cbn; intros u; pose (T_inO u). refine (contr_equiv _ f); exact _. - Defined. - - Definition hprop_inO `{Funext} - (O : Modality@{u a}) (T : Type@{i}) - : IsHProp (In@{u a i} O T). - Proof. + - intros ; exact (Join U X). + - intros T u. + pose (contr_inhabited_hprop _ u). exact _. - Defined. - - Definition O_ind_internal (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - : let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - (forall a, B (to@{u a i} O A a)) -> forall a, B a. - Proof. - simpl; intros f z. - simple refine (Pushout_ind@{i a i j i} B _ _ _ z). - - intros u; apply center, B_inO, u. - - intros a; apply f. - - intros [u a]. + - intros T x. + exact (push (inr x)). + - intros A B B_inO f z. + srefine (Pushout_ind B _ _ _ z). + + intros u; apply center, B_inO, u. + + intros a; apply f. + + intros [u a]. pose (B_inO (push (inr a)) u). apply path_contr. - Defined. - - Definition O_ind_beta_internal (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - (f : forall a : A, B (to O A a)) (a : A) - : O_ind_internal O A B B_inO f (to O A a) = f a - := 1. - - Definition minO_paths (O : Modality@{u a}) (A : Type@{i}) - (A_inO : In@{u a i} O A) - (z z' : A) - : In@{u a i} O (z = z'). - Proof. - intros u; pose (A_inO u); apply contr_paths_contr. - Defined. - - Definition IsSepFor@{u a} (O' O : Modality@{u a}) : Type@{u} - := Empty. - - Definition inO_paths_from_inSepO@{u a i iplus} - (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A) - : In@{u a i} O (x = y) - := Empty_rec sep. - - Definition inSepO_from_inO_paths@{u a i iplus} - (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (op : forall (x y : A), In@{u a i} O (x = y)) - : In@{u a i} O' A - := Empty_rec sep. - -End ClosedModalities. - -Module Import ClM := Modalities_Theory ClosedModalities. -Export ClM.Coercions. -Export ClM.RSU.Coercions. - -Coercion Closed_Modality_to_Modality := - idmap : Closed_Modality -> ClosedModalities.Modality. + - intros; reflexivity. + - intros A A_inO z z' u. + pose (A_inO u). + apply contr_paths_contr. +Defined. (** The closed modality is accessible. *) -Module Accessible_ClosedModalities - <: Accessible_Modalities ClosedModalities. - - Module Os_Theory := Modalities_Theory ClosedModalities. - - Definition acc_gen - := fun (O : ClosedModalities.Modality@{u a}) => - Build_NullGenerators@{a} (unCl O) (fun _ => Empty). - - Definition inO_iff_isnull@{u a i} - (O : ClosedModalities.Modality@{u a}) (X : Type@{i}) - : iff@{i i i} - (ClosedModalities.In@{u a i} O X) - (IsNull (acc_gen O) X). - Proof. - split. - - intros X_inO u. +Global Instance accmodality_closed (U : hProp) + : IsAccModality (Cl U). +Proof. + unshelve econstructor. + - econstructor. + exact (fun _:U => Empty). + - intros X; split. + + intros X_inO u. pose (X_inO u). - apply ooextendable_contr@{a a i i}; exact _. - - intros ext u. + apply ooextendable_contr; exact _. + + intros ext u. exists ((fst (ext u 1%nat) Empty_rec).1 tt); intros x. unfold const in ext. exact ((fst (snd (ext u 2) (fst (ext u 1%nat) Empty_rec).1 (fun _ => x)) (Empty_ind _)).1 tt). - Defined. - -End Accessible_ClosedModalities. +Defined. (** In fact, it is topological, and therefore (assuming univalence) lex. As for topological modalities generally, we don't need to declare these as global instances, but we prove them here as local instances for exposition. *) -Module Import ClT := - Topological_Modalities_Theory - ClosedModalities - Accessible_ClosedModalities. - -Local Instance topological_closed (O : Modality) -: Topological O. -Proof. - exact _. -Defined. +Local Instance topological_closed (U : hProp) + : Topological (Cl U) + := _. -Local Instance lex_closed `{Univalence} (O : Modality) -: Lex O. +Global Instance lex_closed `{Univalence} (U : hProp) + : Lex (Cl U). Proof. - exact _. + rapply lex_topological. Defined. (** Thus, it also has the following alternative version. *) -Definition Cl' (U : hProp) -: Nullification_Modality +Definition Cl' (U : hProp) : Modality := Nul (Build_NullGenerators U (fun _ => Empty)). diff --git a/theories/Modalities/Descent.v b/theories/Modalities/Descent.v new file mode 100644 index 00000000000..40afc580715 --- /dev/null +++ b/theories/Modalities/Descent.v @@ -0,0 +1,441 @@ +(* -*- mode: coq; mode: visual-line -*- *) +Require Import HoTT.Basics HoTT.Types HoTT.Cubical. +Require Import EquivalenceVarieties Fibrations Extensions Pullback NullHomotopy Factorization UnivalenceImpliesFunext PathAny. +Require Import ReflectiveSubuniverse Modality Accessible Localization. + +Local Open Scope path_scope. +Local Open Scope subuniverse_scope. + +(** * Descent between subuniverses *) + +(** We study here a strengthening of the relation [O << O'] saying that [O]-modal type families descend along [O']-equivalences. Pairs of reflective subuniverses with this relation share nearly all the properties of a reflective subuniverse [O] paired with its subuniverse [Sep O] of separated types (see [Separated.v]) and also many of those of a single left exact modality (see [Lex.v]). Thus, many of the results herein generalize those of RSS for lex modalities and those of CORS for separated subuniverses. + +Note that this kind of descent is not the same as the "modal descent" of Cherubini and Rijke. When we get around to formalizing that, we may need to worry about disambiguating the names. *) + +(** ** Definitions *) + +(** This definition is an analogue of the statement of Lemma 2.19 of CORS, and of Theorem 3.1(xiii) of RSS. Note that CORS Lemma 2.19 includes uniqueness of the extension, which we don't assert explicitly. However, uniqueness follows from the [ReflectsD] parameter -- see [ooextendable_TypeO_lex_leq] below. *) +Class Descends@{i} (O' O : Subuniverse@{i}) (T : Type@{i}) + `{ReflectsD@{i} O' O T} := +{ + OO_descend : + forall (P : T -> Type@{i}) {P_inO : forall x, In O (P x)}, + O_reflector O' T -> Type@{i} ; + OO_descend_inO : + forall (P : T -> Type@{i}) {P_inO : forall x, In O (P x)} (x : O_reflector O' T), + In O (OO_descend P x) ; + OO_descend_beta : + forall (P : T -> Type@{i}) {P_inO : forall x, In O (P x)} (x : T), + OO_descend P (to O' T x) <~> P x ; +}. + +Global Existing Instance OO_descend_inO. +Arguments OO_descend O' O {T _ _ _} P {P_inO} x. +Arguments OO_descend_inO O' O {T _ _ _} P {P_inO} x. +Arguments OO_descend_beta O' O {T _ _ _} P {P_inO} x. + +Class O_lex_leq (O1 O2 : ReflectiveSubuniverse) `{O1 << O2} := + O_lex_leq_descends : forall A, Descends O2 O1 A. + +Notation "O1 <<< O2" := (O_lex_leq O1 O2) (at level 70) : subuniverse_scope. +Global Existing Instance O_lex_leq_descends. + +(** Unfortunately, it seems that generalizing binders don't work on notations: writing [`{O <<< O'}] doesn't automatically add the precondition [O << O'], although writing [`{O_lex_leq O O'}] does. *) + +Definition O_lex_leq_eq {O1 O2 O3 : ReflectiveSubuniverse} + `{O1 <=> O2} `{O2 << O3, O2 <<< O3} + (Hstrong := O_strong_leq_trans_l O1 O2 O3) + : O1 <<< O3. +Proof. + intros A; unshelve econstructor; intros P P_inO1. + all:pose (P_inO2 := fun x => inO_leq O1 O2 _ (P_inO1 x)). + - apply (OO_descend O3 O2 P). + - intros x; apply (inO_leq O2 O1), (OO_descend_inO O3 O2 P). + - apply (OO_descend_beta O3 O2 P). +Defined. + +(** ** Left exactness properties *) + +(** We prove analogues of the properties in section 2.4 of CORS and Theorem 3.1 of RSS, but in a different order, with different proofs, to increase the generality. The proofs in CORS use Proposition 2.26 for everything else, but it seems that most of the other results are true in the generality of two reflective subuniverses with [O <<< O'], so we give different proofs for some of them. (To show that this generality is non-spurious, note that a lex modality [O] satisfies [O <<< O], but does not generally coincide with [Sep O].) + +In the case of a single modality, most of these statements are equivalent to lex-ness (as stated in Theorem 3.1 of RSS). We do not know if anything similar is true more generally. *) + +Section LeftExactness. +Context (O' O : ReflectiveSubuniverse) `{O << O', O <<< O'}. + +(** Proposition 2.30 of CORS and Theorem 3.1(xii) of RSS: any [O']-equivalence is [O]-connected. The special case when [f = to O' A] requires only [O << O'], but the general case seems to require [O <<< O']. *) +Global Instance conn_map_OO_inverts + {A B : Type} (f : A -> B) `{O_inverts O' f} + : IsConnMap O f. +Proof. + apply conn_map_from_extension_elim. + intros P P_inO. + assert (E : ExtendableAlong 1%nat f P); [ | exact (fst E) ]. + assert (Qp := OO_descend_beta O' O P). + assert (Q_inO := OO_descend_inO O' O P). + set (Q := OO_descend O' O P) in *. + refine (extendable_postcompose' _ (Q o to O' B) P f Qp _). + refine (cancelL_extendable _ Q f (to O' B) _ _). + 1:srapply (extendable_conn_map_inO O). + refine (extendable_homotopic _ _ (O_functor O' f o to O' A) (to_O_natural O' f) _). + srapply extendable_compose. + 1:srapply extendable_equiv. + srapply (extendable_conn_map_inO O). +Defined. + +(** A generalization of Lemma 2.27 of CORS: [functor_sigma] of a family of [O]-equivalences over an [O']-equivalence is an [O]-equivalence. CORS Lemma 2.27 is the case when [f = to O' A] and [g] is a family of identities. *) +Definition OO_inverts_functor_sigma + {A B : Type} {P : A -> Type} {Q : B -> Type} + (f : A -> B) (g : forall a, P a -> Q (f a)) + `{O_inverts O' f} `{forall a, O_inverts O (g a)} + : O_inverts O (functor_sigma f g). +Proof. + srapply isequiv_homotopic'. + - refine (equiv_O_sigma_O O _ oE _ oE (equiv_O_sigma_O O _)^-1). + refine (Build_Equiv _ _ (O_functor O (functor_sigma f (fun x => O_functor O (g x)))) _). + - apply O_indpaths. intros [x u]; cbn. + rewrite O_rec_beta, to_O_natural, O_rec_beta; cbn. + rewrite !to_O_natural, O_rec_beta. + reflexivity. +Defined. + +(** Families of [O]-modal types descend along all [O']-equivalences (not just the [O']-units, as asserted in the definition of [<<<]. *) +Definition OO_descend_O_inverts + {A B : Type} (f : A -> B) `{O_inverts O' f} + (P : A -> Type) {P_inO : forall x, In O (P x)} + : B -> Type. +Proof. + intros b. + pose (Q := OO_descend O' O P). + exact (Q ((O_functor O' f)^-1 (to O' B b))). +Defined. + +Global Instance OO_descend_O_inverts_inO + {A B : Type} (f : A -> B) `{O_inverts O' f} + (P : A -> Type) {P_inO : forall x, In O (P x)} (b : B) + : In O (OO_descend_O_inverts f P b) + := _. + +Definition OO_descend_O_inverts_beta + {A B : Type} (f : A -> B) `{O_inverts O' f} + (P : A -> Type) {P_inO : forall x, In O (P x)} (a : A) + : (OO_descend_O_inverts f P (f a)) <~> P a. +Proof. + unfold OO_descend_O_inverts. + refine (OO_descend_beta O' O P a oE _). + assert (p := (to_O_natural O' f a)^). + apply moveR_equiv_V in p. + exact (equiv_transport _ _ _ p). +Defined. + +(** Morally, an equivalent way of saying [O <<< O'] is that the universe of [O]-modal types is [O']-modal. We can't say this directly since this type lives in a higher universe, but here is a rephrasing of it. *) +Definition ooextendable_TypeO_lex_leq `{Univalence} + {A B : Type} (f : A -> B) `{O_inverts O' f} + : ooExtendableAlong f (fun _ => Type_ O). +Proof. + (** It suffices to show that maps into [Type_ O] extend along [f], and that sections of families of equivalences are [ooExtendableAlong] it. However, due to the implementation of [ooExtendableAlong], we have to prove the first one twice (there should probably be a general cofixpoint lemma for this). *) + intros [|[|n]]; + [ exact tt + | split; [ intros P | intros; exact tt ] + | split; [ intros P | ] ]. + (** The first follows from what we just proved. *) + 1-2:exists (fun x => (OO_descend_O_inverts f P x ; + OO_descend_O_inverts_inO f P x)). + 1-2:intros x; apply path_TypeO, path_universe_uncurried; cbn. + 1-2:exact (OO_descend_O_inverts_beta f P x). + (** The second follows from the fact that the type of equivalences between two [O]-modal types is [O]-modal. *) + intros P Q; rapply (ooextendable_postcompose' (fun b => P b <~> Q b)). + - intros x; refine (equiv_path_TypeO _ _ _ oE equiv_path_universe _ _). + - (** Typeclass inference spins on [rapply], I don't know why. *) + apply (ooextendable_conn_map_inO O); exact _. +Defined. + +(** We can also state it in terms of belonging to a subuniverse if we lift [O'] accessibly (an analogue of Theorem 3.11(iii) of RSS). *) +Global Instance inO_TypeO_lex_leq `{Univalence} `{IsAccRSU O'} + : In (lift_accrsu O') (Type_ O) + := fun i => ooextendable_TypeO_lex_leq (acc_lgen O' i). + +(** If [f] is an [O']-equivalence, then [ap f] is an [O]-equivalence. *) +Global Instance OO_inverts_ap + {A B : Type} (f : A -> B) `{O_inverts O' f} (x y : A) + : O_inverts O (@ap _ _ f x y). +Proof. + assert (Pb := OO_descend_O_inverts_beta f (fun y:A => O (x = y))). + assert (P_inO := OO_descend_O_inverts_inO f (fun y:A => O (x = y))). + set (P := OO_descend_O_inverts f (fun y:A => O (x = y))) in *. + clearbody P; cbn in *. + srapply isequiv_adjointify. + - intros q. + pose (t := fun p => @transport B P (f x) (f y) p ((Pb x)^-1 (to O (x = x) 1))). + exact (Pb y (O_rec t q)). + - apply O_indpaths; intros p; cbn. + rewrite O_rec_beta. + assert (g := extension_conn_map_elim O (functor_sigma f (fun (a:A) (p:P (f a)) => p)) + (fun bp => O (f x = bp.1)) (fun u => O_functor O (ap f) (Pb u.1 u.2))). + pose (g1 b p := g.1 (b;p)). cbn in g1. + assert (e : (fun u => g1 u.1 u.2) == g.1). + 1:intros [a b]; reflexivity. + assert (g2 := fun a p => e _ @ g.2 (a;p)); cbn in g2. + refine ((g2 y _)^ @ _). + rewrite (ap_transport p g1). + rewrite (g2 x ((Pb x)^-1 (to O (x = x) 1))). + rewrite eisretr, to_O_natural; cbn. + rewrite <- (ap_transport p (fun b => to O (f x = b))). + apply ap. + rewrite transport_paths_r. + apply concat_1p. + - apply O_indpaths; intros p; cbn. + rewrite to_O_natural, O_rec_beta. + destruct p; cbn. + srapply eisretr. +Defined. + +Definition equiv_O_functor_ap_OO_inverts + {A B : Type} (f : A -> B) `{O_inverts O' f} (x y : A) + : O (x = y) <~> O (f x = f y) + := Build_Equiv _ _ (O_functor O (ap f)) _. + +(** Theorem 3.1(i) of RSS: path-spaces of [O']-connected types are [O]-connected. *) +Definition OO_isconnected_paths + {A : Type} `{IsConnected O' A} (x y : A) + : IsConnected O (x = y). +Proof. + rapply (contr_equiv' _ (equiv_O_functor_ap_OO_inverts (const tt) x y)^-1). +Defined. + +(** Proposition 2.26 of CORS and Theorem 3.1(ix) of RSS; also generalizes Theorem 7.3.12 of the book. Here we need to add the extra assumption that [O' <= Sep O], which is satisfied when [O' = Sep O] but also when [O] is lex and [O' = O]. That some such extra hypothesis is necessary can be seen from the fact that [Tr (-2) <<< O'] for any [O'], whereas this statement is certainly not true in that generality. *) +Definition path_OO `{O' <= Sep O} + {X : Type@{i}} (x y : X) + : O (x = y) -> (to O' X x = to O' X y). +Proof. + nrefine (O_rec (O := O) (@ap X (O' X) (to O' X) x y)). + - rapply (@inO_leq O' (Sep O)). + - exact _. +Defined. + +Global Instance isequiv_path_OO `{O' <= Sep O} + {X : Type@{i}} (x y : X) + : IsEquiv (path_OO x y). +Proof. + nrefine (isequiv_O_rec_O_inverts O _). + (** Typeclass search can find this, but it's quicker (and may help the reader) to give it explicitly. *) + apply (OO_inverts_ap (to O' X)). +Defined. + +Definition equiv_path_OO `{O' <= Sep O} + {X : Type@{i}} (x y : X) + : O (x = y) <~> (to O' X x = to O' X y) + := Build_Equiv _ _ (path_OO x y) _. + +(** [functor_hfiber] on a pair of [O']-equivalences is an [O]-equivalence. *) +Global Instance OO_inverts_functor_hfiber + {A B C D : Type} {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + (p : k o f == g o h) (b : B) + `{O_inverts O' h, O_inverts O' k} + : O_inverts O (functor_hfiber p b). +Proof. + unfold functor_hfiber. + simple notypeclasses refine (OO_inverts_functor_sigma _ _). + 1:exact _. + intros a; cbn. + refine (isequiv_homotopic (O_functor O (concat (p a)^) o O_functor O (@ap _ _ k (f a) b)) _). + symmetry; apply O_functor_compose. +Defined. + +(** Corollary 2.29 of CORS: [O'] preserves fibers up to [O]-equivalence. *) +Global Instance OO_inverts_functor_hfiber_to_O + {Y X : Type} (f : Y -> X) (x : X) + : O_inverts O (functor_hfiber (fun a => (to_O_natural O' f a)^) x). +Proof. + (** Typeclass search can find this, but it's faster to give it explicitly. *) + exact (OO_inverts_functor_hfiber _ _). +Defined. + +Definition equiv_OO_functor_hfiber_to_O + {Y X : Type@{i} } (f : Y -> X) (x : X) + : O (hfiber f x) <~> O (hfiber (O_functor O' f) (to O' X x)) + := Build_Equiv _ _ _ (OO_inverts_functor_hfiber_to_O f x). + +(** Theorem 3.1(iii) of RSS: any map between [O']-connected types is [O]-connected. (Part (ii) is just the version for dependent projections.) *) +Definition OO_conn_map_isconnected + {Y X : Type} `{IsConnected O' Y, IsConnected O' X} (f : Y -> X) + : IsConnMap O f. +Proof. + intros x; rapply (contr_equiv' _ (equiv_OO_functor_hfiber_to_O f x)^-1). +Defined. + +Definition OO_isconnected_hfiber + {Y X : Type} `{IsConnected O' Y, IsConnected O' X} (f : Y -> X) (x : X) + : IsConnected O (hfiber f x) + := OO_conn_map_isconnected f x. + +(** Theorem 3.1(iv) of RSS: an [O]-modal map between [O']-connected types is an equivalence. *) +Global Instance OO_isequiv_mapino_isconnected + {Y X : Type} `{IsConnected O' Y, IsConnected O' X} (f : Y -> X) `{MapIn O _ _ f} + : IsEquiv f. +Proof. + apply (isequiv_conn_ino_map O). + - apply OO_conn_map_isconnected. + - assumption. +Defined. + +(** Theorem 3.1(vi) of RSS (and part (v) is just the analogue for dependent projections). *) +Definition OO_conn_map_functor_hfiber {A B C D : Type} + {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + `{IsConnMap O' _ _ h, IsConnMap O' _ _ k} + (p : k o f == g o h) (b : B) + : IsConnMap O (functor_hfiber p b). +Proof. + intros [c q]. + nrefine (isconnected_equiv' O _ (hfiber_functor_hfiber p b c q)^-1 _). + apply OO_isconnected_hfiber. +Defined. + +(** An enhancement of Corollary 2.29 of CORS, corresponding to Theorem 3.1(viii) of RSS: when [O'] is a modality, the map between fibers is not just an O-equivalence but is O-connected. *) +Global Instance OO_conn_map_functor_hfiber_to_O `{IsModality O'} + {Y X : Type} (f : Y -> X) (x : X) + : IsConnMap O (functor_hfiber (fun y => (to_O_natural O' f y)^) x). +Proof. + apply OO_conn_map_functor_hfiber. +Defined. + +(** Theorem 3.1(vii) of RSS *) +Definition OO_ispullback_connmap_mapino + {A B C D : Type} {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + (p : k o f == g o h) + `{O_inverts O' h, O_inverts O' k, MapIn O _ _ f, MapIn O _ _ g} + : IsPullback p. +Proof. + apply ispullback_isequiv_functor_hfiber; intros b. + apply (isequiv_O_inverts O). + apply OO_inverts_functor_hfiber; exact _. +Defined. + +(** [functor_pullback] on a triple of [O']-equivalences is an [O]-equivalence. *) +Global Instance OO_inverts_functor_pullback + {A1 B1 C1 A2 B2 C2 : Type} + (f1 : B1 -> A1) (g1 : C1 -> A1) + (f2 : B2 -> A2) (g2 : C2 -> A2) + (h : A1 -> A2) (k : B1 -> B2) (l : C1 -> C2) + (p : f2 o k == h o f1) (q : g2 o l == h o g1) + `{O_inverts O' h, O_inverts O' k, O_inverts O' l} + : O_inverts O (functor_pullback f1 g1 f2 g2 h k l p q). +Proof. + unfold functor_pullback. + simple notypeclasses refine (OO_inverts_functor_sigma _ _). + 1:exact _. + intros b1; cbn. + simple notypeclasses refine (OO_inverts_functor_sigma _ _). + 1:exact _. + intros c1; cbn. + refine (isequiv_homotopic (O_functor O (fun r => r @ (q c1)^) o O_functor O (concat (p b1)) o O_functor O (@ap _ _ h (f1 b1) (g1 c1))) _). + intros r; symmetry. + refine (_ @ _). + 2:apply O_functor_compose. + cbn; srapply O_functor_compose. +Defined. + +(** Proposition 2.28 of CORS, and Theorem 3.1(x) of RSS: the functor [O'] preserves pullbacks up to [O]-equivalence. *) +Global Instance OO_inverts_functor_pullback_to_O + {A B C : Type} (f : B -> A) (g : C -> A) + : O_inverts O (functor_pullback f g (O_functor O' f) (O_functor O' g) + (to O' A) (to O' B) (to O' C) + (to_O_natural O' f) (to_O_natural O' g)). +Proof. + apply OO_inverts_functor_pullback; exact _. +Defined. + +Definition equiv_OO_pullback {A B C : Type} (f : B -> A) (g : C -> A) + : O (Pullback f g) <~> O (Pullback (O_functor O' f) (O_functor O' g)) + := Build_Equiv _ _ _ (OO_inverts_functor_pullback_to_O f g). + +(** The "if" direction of CORS Proposition 2.31, and the nontrivial part of Theorem 3.1(xi) of RSS. Note that we could also deduce Theorem 3.1(iii) of RSS from this. *) +Definition OO_cancelL_conn_map + {Y X Z : Type} (f : Y -> X) (g : X -> Z) + `{IsConnMap O' _ _ (g o f)} `{IsConnMap O' _ _ g} + : IsConnMap O f. +Proof. + apply conn_map_OO_inverts. + nrapply (cancelL_isequiv (O_functor O' g)). + 1:exact _. + rapply (isequiv_homotopic _ (O_functor_compose O' f g)). +Defined. + +End LeftExactness. + +(** Here's the "only if" direction of CORS Proposition 2.31. Note that the hypotheses are different from those of the "if" direction, and the proof is shorter than the one given in CORS. *) +Definition OO_cancelR_conn_map + (O' O : ReflectiveSubuniverse) `{O <= O', O' <= Sep O} + {Y X Z : Type} (f : Y -> X) (g : X -> Z) + `{IsConnMap O' _ _ (g o f)} `{IsConnMap O _ _ f} + : IsConnMap O' g. +Proof. + apply conn_map_from_extension_elim. + intros P P_inO h. + exists (conn_map_elim O' (g o f) P (h o f)). + refine (conn_map_elim O f _ _). + - intros x. + pose proof (fun z => inO_leq O' (Sep O) (P z) (P_inO z)). + exact _. + - intros y. + apply (conn_map_comp O' (g o f)). +Defined. + +Definition OO_isconnected_from_conn_map + (O' O : ReflectiveSubuniverse) `{O <= O', O' <= Sep O} + {Y X : Type} (f : Y -> X) + `{IsConnected O' Y} `{IsConnMap O _ _ f} + : IsConnected O' X. +Proof. + apply isconnected_conn_map_to_unit. + apply (OO_cancelR_conn_map O' O f (const tt)). +Defined. + +(** An interesting scholium to Proposition 2.31. *) +Definition OO_inverts_conn_map_factor_conn_map + (O' O : ReflectiveSubuniverse) `{O << O', O <<< O', O' <= Sep O} + {Y X Z : Type} (f : Y -> X) (g : X -> Z) + `{IsConnMap O' _ _ (g o f)} `{IsConnMap O _ _ f} + : O_inverts O' f. +Proof. + nrapply (cancelL_isequiv (O_functor O' g)). + - apply O_inverts_conn_map. + apply (OO_cancelR_conn_map O' O f g). + - rapply (isequiv_homotopic _ (O_functor_compose O' f g)). +Defined. + +Definition OO_inverts_conn_map_isconnected_domain + (O' O : ReflectiveSubuniverse) `{O << O', O <<< O', O' <= Sep O} + {Y X : Type} (f : Y -> X) + `{IsConnected O' Y} `{IsConnMap O _ _ f} + : O_inverts O' f. +Proof. + apply (OO_inverts_conn_map_factor_conn_map O' O f (const tt)). +Defined. + +(** Here is the converse of [ooextendable_TypeO_lex_leq]. *) +Definition O_lex_leq_extendable_TypeO + (O' O : ReflectiveSubuniverse) `{O << O'} + (e : forall (A:Type) (g:A->Type_ O), ExtensionAlong (to O' A) (fun _ => Type_ O) g) + : O <<< O'. +Proof. + intros A; unshelve econstructor; intros P' P_inO; pose (P := fun x => (P' x ; P_inO x) : Type_ O). + - exact (fun x => ((e A P).1 x).1). + - exact (fun x => ((e A P).1 x).2). + - intros x. + apply equiv_path. + exact (((e A P).2 x)..1). +Defined. + +(** And a version for the accessible case. *) +Definition O_lex_leq_inO_TypeO + (O' O : ReflectiveSubuniverse) `{O << O'} + `{IsAccRSU O'} `{In (lift_accrsu O') (Type_ O)} + : O <<< O'. +Proof. + apply O_lex_leq_extendable_TypeO. + intros A g. + assert (O_inverts (lift_accrsu O') (to O' A)). + - rapply (O_inverts_O_leq' (lift_accrsu O') O'). + - exact (fst (ooextendable_O_inverts (lift_accrsu O') (to O' A) (Type_ O) 1%nat) g). +Defined. diff --git a/theories/Modalities/Fracture.v b/theories/Modalities/Fracture.v index 209e7631a2f..f2d132fb78b 100644 --- a/theories/Modalities/Fracture.v +++ b/theories/Modalities/Fracture.v @@ -25,10 +25,6 @@ We will prove the fracture theorem holds under the assumptions that [O2] is lex, It may sometimes happen that in addition, the "intersection" of [O1] and [O2] is trivial. This is naturally expressed in the context of the fracture theorem by saying that [O2]-modal types are [O1]-connected, i.e. the converse of the second hypothesis of our fracture theorem. When this also holds, we can show that the universe [Type] can actually be reconstructed, up to equivalence, from the universes of [O1]- and [O2]-modal types and the [O2]-reflection from the first to the second, using the "Artin gluing construction" from topos theory. *) -Module Fracture (Os : Modalities). - - Module Export OsL := Lex_Modalities_Theory Os. - (** ** The fracture theorem *) Section FractureTheorem. @@ -49,8 +45,10 @@ Module Fracture (Os : Modalities). Definition ispullback_fracture_square A : IsPullback (fracture_square A). Proof. - refine (ispullback_connmap_mapino_commsq O2 _). - (** And typeclass magic finds everything except the one we need our hypothesis for. *) + apply ispullback_symm. + nrefine (ispullback_connmap_mapino_commsq O2 _). + 1-3:exact _. + 2:rapply mapinO_between_inO. intros x; refine (ino2_isconnectedo1 _ _). Defined. @@ -156,17 +154,9 @@ Module Fracture (Os : Modalities). End FractureTheorem. -End Fracture. - (** ** The propositional fracture theorem *) -(** An easy example of the lex-modal fracture theorem is supplied by the open and closed modalities for an hprop [U]. In order to consider these together, we need to build their union as a family of modalities. *) - -Module Open_And_Closed_Modalities - := Modalities_FamUnion OpenModalities ClosedModalities. -Module Import OpClM := Lex_Modalities_Theory Open_And_Closed_Modalities. -Module Import Propositional_Fracture - := Fracture Open_And_Closed_Modalities. +(** An easy example of the lex-modal fracture theorem is supplied by the open and closed modalities for an hprop [U]. *) Definition gluable_open_closed `{Funext} (U : hProp) : Gluable (Op U) (Cl U). @@ -189,10 +179,6 @@ Defined. (** We can also prove the same thing without funext if we use the nullification versions of these modalities. *) -Import NulM. -Module Import Propositional_Fracture' - := Fracture Nullification_Modalities. - Definition gluable_open_closed' (U : hProp) : Gluable (Op' U) (Cl' U). Proof. @@ -207,7 +193,7 @@ Proof. - refine (isequiv_adjointify _ (to (Op' U) A) _ _). + intros a; apply O_rec_beta. + intros oa; revert oa; apply O_indpaths; intros a; simpl. - apply ap, O_rec_beta. } + apply ap. rapply O_rec_beta. } apply ooextendable_contr; exact _. Defined. diff --git a/theories/Modalities/Identity.v b/theories/Modalities/Identity.v index 2a76033c5f0..067a2360d69 100644 --- a/theories/Modalities/Identity.v +++ b/theories/Modalities/Identity.v @@ -9,106 +9,21 @@ Local Open Scope path_scope. (** Everything to say here is fairly trivial. *) -Inductive Identity_Modality : Type1 - := purely : Identity_Modality. - -Module Identity_Modalities <: Modalities. - - Definition Modality : Type2@{u a} - := Identity_Modality@{a}. - - Definition O_reflector : forall (O : Modality@{u a}), - Type@{i} -> Type@{i} - := fun O X => X. - - Definition In : forall (O : Modality@{u a}), - Type@{i} -> Type@{i} - := fun O X => Unit. - - Definition O_inO : forall (O : Modality@{u a}) (T : Type@{i}), - In@{u a i} O (O_reflector@{u a i} O T) - := fun O X => tt. - - Definition to : forall (O : Modality@{u a}) (T : Type@{i}), - T -> O_reflector@{u a i} O T - := fun O X x => x. - - Definition inO_equiv_inO : - forall (O : Modality@{u a}) (T : Type@{i}) (U : Type@{j}) - (T_inO : In@{u a i} O T) (f : T -> U) (feq : IsEquiv f), - let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - In@{u a j} O U - := fun O T U _ _ _ => tt. - - Definition hprop_inO@{u a i} - : Funext -> forall (O : Modality@{u a}) (T : Type@{i}), - IsHProp (In@{u a i} O T) - := fun _ O T => trunc_contr@{i}. - - Definition O_ind_internal - : forall (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)), - let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - (forall a, B (to O A a)) -> forall a, B a - := fun O A B _ f a => f a. - - Definition O_ind_beta_internal - : forall (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - (f : forall a : A, B (to O A a)) (a:A), - O_ind_internal O A B B_inO f (to O A a) = f a - := fun _ _ _ _ _ _ => 1. - - Definition minO_paths - : forall (O : Modality@{u a}) - (A : Type@{i}) (A_inO : In@{u a i} O A) (z z' : A), - In@{u a i} O (z = z') - := fun _ _ _ _ _ => tt. - - Definition IsSepFor@{u a} - : forall (O' O : Modality@{u a}), Type@{u} - := fun _ _ => Unit. - - Definition inO_paths_from_inSepO@{u a i iplus} - : forall (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A), - In@{u a i} O (x = y) - := fun _ _ _ _ _ _ _ => tt. - - Definition inSepO_from_inO_paths@{u a i iplus} - : forall (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}), - (forall (x y : A), In@{u a i} O (x = y)) -> In@{u a i} O' A - := fun _ _ _ _ _ => tt. - -End Identity_Modalities. - -Module purelyM := Modalities_Theory Identity_Modalities. -Export purelyM.Coercions. -Export purelyM.RSU.Coercions. - -Coercion Identity_Modalities_to_Modalities := idmap - : Identity_Modality -> Identity_Modalities.Modality. - - -Module Accessible_Identity <: Accessible_Modalities Identity_Modalities. - - Module Import Os_Theory := Modalities_Theory Identity_Modalities. - - Definition acc_gen : Modality@{u a} -> NullGenerators@{a} - := fun _ => Build_NullGenerators Empty (fun _ => Empty). - - Definition inO_iff_isnull@{u a i} - : forall (O : Modality@{u a}) (X : Type@{i}), - iff@{i i i} - (In@{u a i} O X) - (IsNull_Internal.IsNull@{a i} (acc_gen O) X) - := fun O X => @pair _ (_ -> Unit) - (fun _ => Empty_ind _) - (fun _ => tt). - -End Accessible_Identity. +Definition purely : Modality. +Proof. + srapply (Build_Modality (fun _ => Unit) _ _ idmap). + 1-2,6:intros; exact tt. + - intros; assumption. + - intros ? ? ? f z; exact (f z). + - intros; reflexivity. +Defined. + +Global Instance accmodality_purely : IsAccModality purely. +Proof. + unshelve econstructor. + - econstructor. + exact (@Empty_rec Type). + - intros X; split. + + intros _ []. + + intros; exact tt. +Defined. diff --git a/theories/Modalities/Lex.v b/theories/Modalities/Lex.v index f73971cc7a2..d4df2d40702 100644 --- a/theories/Modalities/Lex.v +++ b/theories/Modalities/Lex.v @@ -1,182 +1,279 @@ (* -*- mode: coq; mode: visual-line -*- *) Require Import HoTT.Basics HoTT.Types. -Require Import EquivalenceVarieties Fibrations Extensions Pullback NullHomotopy Factorization UnivalenceImpliesFunext. -Require Import Modality Accessible. +Require Import EquivalenceVarieties Fibrations Extensions Pullback NullHomotopy Factorization UnivalenceImpliesFunext PathAny. +Require Import Modality Accessible Localization Descent Separated. Require Import HoTT.Tactics. Local Open Scope path_scope. - +Local Open Scope subuniverse_scope. (** * Lex modalities *) -(** ** Basic theory *) +(** A lex modality is one that preserves finite limits, or equivalently pullbacks. Many equivalent characterizations of this can be found in Theorem 3.1 of RSS. -(** A lex modality is one that preserves finite limits, or equivalently pullbacks. It turns out that a more basic and useful way to say this is that all path-spaces of connected types are connected. Note how different this is from the behavior of, say, truncation modalities! +We choose as our definition that a lex modality to be a reflective subuniverse such that [O <<< O], which is close to (but not quite the same as) RSS Theorem 3.1 (xiii). - This is a "large" definition, and we don't know of any small one that's equivalent to it (see . However, so far we never need to apply it "at multiple universes at once". Thus, rather than making it a module type, we can make it a typeclass and rely on ordinary universe polymorphism. *) +Note that since this includes [O << O] as a precondition, such an [O] must indeed be a modality (and since modalities coerce to reflective subuniverses, in the following notation [O] could be either an element of [ReflectiveSubuniverse] or of [Modality]). *) +Notation Lex O := (O <<< O). -Module Lex_Modalities_Theory (Os : Modalities). +(** ** Properties of lex modalities *) - Module Export Os_Theory := Modalities_Theory Os. +(** We now show that lex modalities have all the other properties from RSS Theorem 3.1 (which are equivalent to lex-ness). All of them are simple specializations of properties from [Descent.v] to the case [O' = O] (although in the general case they are not known to be equivalent). *) +Section LexModality. + Context (O : Modality) `{Lex O}. - Class Lex (O : Modality@{u a}) - := isconnected_paths : forall (A : Type@{i}) (x y : A), - IsConnected@{u a i} O A -> - IsConnected@{u a i} O (x = y). + (** RSS Theorem 3.1 (i) *) + Definition isconnected_paths + {A : Type} `{IsConnected O A} (x y : A) + : IsConnected O (x = y) + := OO_isconnected_paths O O x y. - Global Existing Instance isconnected_paths. + (** RSS Theorem 3.1 (iii) *) + Definition conn_map_lex + {Y X : Type} `{IsConnected O Y, IsConnected O X} (f : Y -> X) + : IsConnMap O f + := OO_conn_map_isconnected O O f. - (** The following numbered lemmas are all actually equivalent characterizations of lex-ness. We prove this for some of them, but we don't make the reverse implications Instances; usually [isconnected_paths] is the easier way to prove lexness. *) + (** RSS Theorem 3.1 (iv) *) + Global Instance isequiv_mapino_isconnected + {Y X : Type} `{IsConnected O Y, IsConnected O X} + (f : Y -> X) `{MapIn O _ _ f} + : IsEquiv f + := OO_isequiv_mapino_isconnected O O f. - (** 1. Every map between connected types is a connected map. *) - Global Instance conn_map_lex {O : Modality} `{Lex O} - {A : Type@{i}} {B : Type@{j}} {f : A -> B} - `{IsConnected O A} `{IsConnected O B} - : IsConnMap O f. + (** RSS Theorem 3.1 (vi) *) + Definition conn_map_functor_hfiber {A B C D : Type} + {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + `{IsConnMap O _ _ h, IsConnMap O _ _ k} + (p : k o f == g o h) (b : B) + : IsConnMap O (functor_hfiber p b) + := OO_conn_map_functor_hfiber O O p b. + + (** RSS Theorem 3.1 (vii) *) + Definition ispullback_connmap_mapino_commsq + {A B C D : Type} {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} + (p : k o f == g o h) + `{O_inverts O h, O_inverts O k, MapIn O _ _ f, MapIn O _ _ g} + : IsPullback p + := OO_ispullback_connmap_mapino O O p. + + (** RSS Theorem 3.1 (viii) *) + Global Instance conn_map_functor_hfiber_to_O + {Y X : Type} (f : Y -> X) (x : X) + : IsConnMap O (functor_hfiber (fun y => (to_O_natural O f y)^) x) + := OO_conn_map_functor_hfiber_to_O O O f x. + + Global Instance isequiv_O_functor_hfiber + {A B} (f : A -> B) (b : B) + : IsEquiv (O_functor_hfiber O f b). Proof. - intros b; refine (isconnected_sigma O). + apply (isequiv_O_rec_O_inverts O). + apply O_inverts_conn_map. + refine (conn_map_homotopic + O (functor_hfiber (fun x => (to_O_natural O f x)^) b) + _ _ _). + intros [a p]. + unfold functor_hfiber, functor_sigma. apply ap. + apply whiskerR, inv_V. Defined. - Definition lex_from_conn_map_lex {O : Modality} - (H : forall A B (f : A -> B), - (IsConnected O A) -> (IsConnected O B) -> - IsConnMap O f) - : Lex O. - Proof. - intros A x y AC. - refine (isconnected_equiv' O (hfiber (unit_name x) y) _ _). - unfold hfiber. - refine (equiv_contr_sigma (fun _ => x = y)). - Defined. + Definition equiv_O_functor_hfiber + {A B} (f : A -> B) (b : B) + : O (hfiber f b) <~> hfiber (O_functor O f) (to O B b) + := Build_Equiv _ _ (O_functor_hfiber O f b) _. - (** 2. Connected maps are left- as well as right-cancellable. *) - Definition cancelL_conn_map (O : Modality) `{Lex O} - {A B C : Type} (f : A -> B) (g : B -> C) - : IsConnMap O g -> IsConnMap O (g o f) -> IsConnMap O f. + (** RSS Theorem 3.1 (ix) *) + Global Instance isequiv_path_O + {X : Type@{i}} (x y : X) + : IsEquiv (path_OO O O x y) + := isequiv_path_OO O O x y. + + Definition equiv_path_O {X : Type@{i}} (x y : X) + : O (x = y) <~> (to O X x = to O X y) + := equiv_path_OO O O x y. + + (** RSS Theorem 3.1 (x). This justifies the term "left exact". *) + Global Instance O_inverts_functor_pullback_to_O + {A B C : Type} (f : B -> A) (g : C -> A) + : O_inverts O (functor_pullback f g (O_functor O f) (O_functor O g) + (to O A) (to O B) (to O C) + (to_O_natural O f) (to_O_natural O g)) + := OO_inverts_functor_pullback_to_O O O f g. + + Definition equiv_O_pullback {A B C : Type} (f : B -> A) (g : C -> A) + : O (Pullback f g) <~> O (Pullback (O_functor O f) (O_functor O g)) + := equiv_OO_pullback O O f g. + + Definition O_functor_pullback + {A B C : Type} (f : B -> A) (g : C -> A) + : IsPullback (O_functor_square O _ _ _ _ (pullback_commsq f g)). Proof. - intros ? ? b. - refine (isconnected_equiv O _ (hfiber_hfiber_compose_map f g b) _). - Defined. + unfold IsPullback. + nrapply (isequiv_homotopic + (O_rec (functor_pullback _ _ _ _ _ _ _ + (to_O_natural O f) (to_O_natural O g)))). + 1:apply isequiv_O_rec_O_inverts; exact _. + apply O_indpaths; intros [b [c e]]. + refine (O_rec_beta _ _ @ _). + (** This *seems* like it ought to be the easier goal, but it turns out to involve lots of naturality wrangling. If we ever want to make real use of this theorem, we might want to separate out this goal into an opaque lemma so we could make the main theorem transparent. *) + unfold functor_pullback, functor_sigma, pullback_corec; simpl. + refine (path_sigma' _ (to_O_natural O pullback_pr1 (b;(c;e)))^ _). + rewrite transport_sigma'; simpl. + refine (path_sigma' _ (to_O_natural O pullback_pr2 (b;(c;e)))^ _). + rewrite transport_paths_Fl. + rewrite transport_paths_Fr. + Open Scope long_path_scope. + unfold O_functor_square. + rewrite ap_V, inv_V, O_functor_homotopy_beta, !concat_p_pp. + unfold pullback_commsq; simpl. + rewrite to_O_natural_compose, !concat_pp_p. + do 3 apply whiskerL. + rewrite ap_V, <- inv_pp. + rewrite <- (inv_V (O_functor_compose _ _ _ _)), <- inv_pp. + apply inverse2, to_O_natural_compose. + Close Scope long_path_scope. + Qed. - (** 3. Every map inverted by [O] is [O]-connected. *) - Definition isconnected_O_inverts (O : Modality) `{Lex O} - {A B : Type} (f : A -> B) `{O_inverts O f} - : IsConnMap O f. + (** RSS Theorem 3.1 (xi) *) + Definition cancelL_conn_map + {Y X Z : Type} (f : Y -> X) (g : X -> Z) + `{IsConnMap O _ _ (g o f)} `{IsConnMap O _ _ g} + : IsConnMap O f + := OO_cancelL_conn_map O O f g. + + (** RSS Theorem 3.1 (xii) *) + Global Instance conn_map_O_inverts + {A B : Type} (f : A -> B) `{O_inverts O f} + : IsConnMap O f + := conn_map_OO_inverts O O f. + + (** RSS Theorem 3.1 (xiii) *) + Definition modal_over_connected_isconst_lex + (A : Type) `{IsConnected O A} + (P : A -> Type) `{forall x, In O (P x)} + : {Q : Type & In O Q * forall x, Q <~> P x}. + Proof. + pose proof (O_inverts_isconnected O (fun _:A => tt)). + exists (OO_descend_O_inverts O O (fun _:A => tt) P tt); split. + - apply OO_descend_O_inverts_inO. + - intros; rapply OO_descend_O_inverts_beta. + Defined. + + (** RSS Theorem 3.11 (iii): in the accessible case, the universe is modal. *) + Global Instance inO_typeO_lex `{Univalence} `{IsAccRSU O} + : In (lift_accrsu O) (Type_ O) + := _. + + (** Part of RSS Corollary 3.9: lex modalities preserve [n]-types for all [n]. This is definitely not equivalent to lex-ness, since it is true for the truncation modalities that are not lex. But it is also not true of all modalities; e.g. the shape modality in a cohesive topos can take 0-types to [oo]-types. With a little more work, this can probably be proven without [Funext]. *) + Global Instance istrunc_O_lex `{Funext} + {n : trunc_index} {A : Type} `{IsTrunc n A} + : IsTrunc n (O A). Proof. - refine (cancelL_conn_map O f (to O B) _ _). - refine (conn_map_homotopic O _ _ (to_O_natural O f) _). - (** Typeclass magic! *) + generalize dependent A; simple_induction n n IHn; intros A ?. + - exact _. (** Already proven for all modalities. *) + - refine (O_ind (fun x => forall y, IsTrunc n (x = y)) _); intros x. + refine (O_ind (fun y => IsTrunc n (to O A x = y)) _); intros y. + refine (trunc_equiv _ (equiv_path_O x y)). Defined. - (** 4. Connected types are closed under pullbacks. (Closure under fibers is [conn_map_lex] above. *) - Global Instance isconnected_pullback (O : Modality) `{Lex O} - {A B C : Type} {f : A -> C} {g : B -> C} - `{IsConnected O A} `{IsConnected O B} `{IsConnected O C} - : IsConnected O (Pullback f g). +End LexModality. + +(** ** Equivalent characterizations of lex-ness *) + +(** We will not prove that *all* of the above properties from RSS Theorem 3.1 are equivalent to lex-ness, but we will do it for some of them. *) + +Section ImpliesLex. + Context {O : Modality}. + + (** RSS 3.1 (xiii) implies lexness *) + Definition lex_from_modal_over_connected_isconst + (H : forall (A : Type) (A_isC : IsConnected O A) + (P : A -> Type) (P_inO : forall x, In O (P x)), + {Q : Type & In O Q * forall x, Q <~> P x}) + : Lex O. Proof. - apply isconnected_sigma; [ exact _ | intros a ]. - refine (isconnected_equiv O (hfiber g (f a)) - (equiv_functor_sigma_id - (fun b => equiv_path_inverse _ _)) - _). + intros A; unshelve econstructor; intros P P_inO. + all:pose (Q := fun z:O A => H (hfiber (to O A) z) _ (P o pr1) _). + - exact (fun z => (Q z).1). + - exact (fun z => fst (Q z).2). + - intros x; cbn. + exact (snd (Q (to O A x)).2 (x;1)). Defined. - (** 5. The reflector preserves pullbacks. This justifies the terminology "lex". *) - Definition O_functor_pullback (O : Modality) `{Lex O} - {A B C} (f : B -> A) (g : C -> A) - : IsPullback (O_functor_square O _ _ _ _ (pullback_commsq f g)). + (** RSS 3.11 (iii), the universe is modal, implies lex-ness *) + Definition lex_from_inO_typeO `{IsAccRSU O} `{In (lift_accrsu O) (Type_ O)} + : Lex O. Proof. - refine (isequiv_O_inverts O _). - refine (O_inverts_conn_map O _). - refine (cancelR_conn_map O (to O (Pullback f g)) _). - refine (conn_map_homotopic O - (functor_pullback f g (O_functor O f) (O_functor O g) - (to O A) (to O B) (to O C) - (to_O_natural O f) (to_O_natural O g)) - _ _ _). - (** This *seems* like it ought to be the easier goal, but it turns out to involve lots of naturality wrangling. If we ever want to make real use of this theorem, we might want to separate out this goal into an opaque lemma so we could make the main theorem transparent. *) - - intros [b [c e]]; - unfold functor_pullback, functor_sigma, pullback_corec; - simpl. - refine (path_sigma' _ (to_O_natural O pullback_pr1 (b;(c;e)))^ _). - rewrite transport_sigma'; simpl. - refine (path_sigma' _ (to_O_natural O pullback_pr2 (b;(c;e)))^ _). - rewrite transport_paths_Fl. - rewrite transport_paths_Fr. - Open Scope long_path_scope. - unfold O_functor_square. - rewrite ap_V, inv_V, O_functor_homotopy_beta, !concat_p_pp. - unfold pullback_commsq; simpl. - rewrite to_O_natural_compose, !concat_pp_p. - do 3 apply whiskerL. - rewrite ap_V, <- inv_pp. - rewrite <- (inv_V (O_functor_compose _ _ _ _)), <- inv_pp. - apply inverse2, to_O_natural_compose. - Close Scope long_path_scope. - (** By contrast, this goal, which seems to contain all the mathematical content, is solved fairly easily by [hfiber_functor_pullback] and typeclass magic invoking [isconnected_pullback]. *) - - intros [ob [oc oe]]. - refine (isconnected_equiv O _ - (hfiber_functor_pullback _ _ _ _ _ _ _ _ _ _)^-1 _). - Qed. + apply (O_lex_leq_inO_TypeO O O). + Defined. - (** 6. The reflector preserves fibers. This is a slightly simpler version of the previous. *) - Global Instance isequiv_O_functor_hfiber (O : Modality) `{Lex O} - {A B} (f : A -> B) (b : B) - : IsEquiv (O_functor_hfiber O f b). + (** RSS Theorem 3.1 (xi) implies lex-ness *) + Definition lex_from_cancelL_conn_map + (cancel : forall {Y X Z : Type} (f : Y -> X) (g : X -> Z), + (IsConnMap O (g o f)) -> (IsConnMap O g) + -> IsConnMap O f) + : Lex O. Proof. - refine (isequiv_O_inverts O _). - apply O_inverts_conn_map. - refine (cancelR_conn_map O (to O _) _). - unfold O_functor_hfiber. - refine (conn_map_homotopic O - (@functor_hfiber _ _ _ _ f (O_functor O f) - (to O A) (to O B) - (fun x => (to_O_natural O f x)^) b) - _ _ _). - - intros [a p]. - rewrite O_rec_beta. - unfold functor_hfiber, functor_sigma. apply ap. - apply whiskerR, inv_V. - - intros [oa p]. - refine (isconnected_equiv O _ - (hfiber_functor_hfiber _ _ _ _)^-1 _). + apply lex_from_modal_over_connected_isconst; intros. + exists (O {x:A & P x}); split; [ exact _ | intros x; symmetry ]. + refine (Build_Equiv _ _ (fun p => to O _ (x ; p)) _). + nrefine (isequiv_conn_map_ino O _). 1-2:exact _. + revert x; apply conn_map_fiber. + nrefine (cancel _ _ _ _ (fun z:{x:A & O {x : A & P x}} => z.2) _ _). + 1: clear cancel; exact _. + intros z. + refine (isconnected_equiv' O A _ _). + unfold hfiber. + refine (equiv_adjointify (fun x => ((x ; z) ; 1)) + (fun y => y.1.1) _ _). + - intros [[x y] []]; reflexivity. + - intros x; reflexivity. Defined. - Definition equiv_O_functor_hfiber (O : Modality) `{Lex O} - {A B} (f : A -> B) (b : B) - : O (hfiber f b) <~> hfiber (O_functor O f) (to O B b) - := Build_Equiv _ _ (O_functor_hfiber O f b) _. - - (** 7. Lex modalities preserve path-spaces. *) - Definition O_path_cmp (O : Modality) {A} (x y : A) - : O (x = y) -> (to O A x = to O A y) - := O_rec (ap (to O A)). - - Global Instance isequiv_O_path_cmp {O : Modality} `{Lex O} {A} (x y : A) - : IsEquiv (O_path_cmp O x y). + (** RSS Theorem 3.1 (iii) implies lex-ness *) + Definition lex_from_conn_map_lex + (H : forall A B (f : A -> B), + (IsConnected O A) -> (IsConnected O B) -> + IsConnMap O f) + : Lex O. Proof. - refine (isequiv_conn_ino_map O _). - refine (cancelR_conn_map O (to O (x = y)) _). - refine (conn_map_homotopic O (ap (to O A)) _ _ _). - - intros ?; symmetry; by apply O_rec_beta. - - intros p. - refine (isconnected_equiv O _ (hfiber_ap p)^-1 _). + apply lex_from_cancelL_conn_map. + intros Y X Z f g gfc gc x. + pose (h := @functor_hfiber Y Z X Z (g o f) g f idmap (fun a => 1%path)). + assert (cc := H _ _ (h (g x)) (gfc (g x)) (gc (g x))). + refine (isconnected_equiv' O _ _ (cc (x;1))). + unfold hfiber. + subst h; unfold functor_hfiber, functor_sigma; cbn. + refine (_ oE (equiv_sigma_assoc _ _)^-1). + apply equiv_functor_sigma_id; intros y; cbn. + refine (_ oE (equiv_functor_sigma_id _)). + 2:intros; symmetry; apply equiv_path_sigma. + cbn. + refine (_ oE equiv_sigma_symm _). + apply equiv_sigma_contr; intros p. + destruct p; cbn. + refine (contr_equiv' { p : g (f y) = g (f y) & p = 1%path } _). + apply equiv_functor_sigma_id; intros p; cbn. + apply equiv_concat_l. + exact (concat_1p _ @ ap_idmap _). Defined. - (** 8. Any modal map between connected types is an equivalence. *) - Global Instance isequiv_ismodal_isconnected_types - {O : Modality} `{Lex O} {A B} {f : A -> B} - `{IsConnected O A} `{IsConnected O B} `{MapIn O _ _ f} - : IsEquiv f. + (** RSS Theorem 3.1 (i) implies lex-ness *) + Definition lex_from_isconnected_paths + (H : forall (A : Type) (Ac : IsConnected O A) (x y : A), + IsConnected O (x = y)) + : Lex O. Proof. - apply (isequiv_conn_ino_map O); exact _. + apply lex_from_conn_map_lex. + intros A B f Ac Bc c. + rapply isconnected_sigma. Defined. + (** RSS Theorem 3.1 (iv) implies lex-ness *) Definition lex_from_isequiv_ismodal_isconnected_types - {O : Modality} (H : forall A B (f : A -> B), - (IsConnected O A) -> (IsConnected O B) -> - (MapIn O f) -> IsEquiv f) + (IsConnected O A) -> (IsConnected O B) -> + (MapIn O f) -> IsEquiv f) : Lex O. Proof. apply lex_from_conn_map_lex. @@ -189,21 +286,8 @@ Module Lex_Modalities_Theory (Os : Modalities). apply (cancelR_conn_map O (factor1 (image O f)) (const tt)). Defined. - (** 9. Any commutative square with connected maps in one direction and modal ones in the other must necessarily be a pullback. *) - Definition ispullback_connmap_mapino_commsq (O : Modality) `{Lex O} {A B C D} - {f : A -> B} {g : C -> D} {h : A -> C} {k : B -> D} - `{IsConnMap O _ _ f} `{IsConnMap O _ _ g} - `{MapIn O _ _ h} `{MapIn O _ _ k} - (p : k o f == g o h) - : IsPullback p. - Proof. - refine (isequiv_conn_ino_map O (pullback_corec p)). - - refine (cancelL_conn_map O (pullback_corec p) (k^* g) _ _). - - refine (cancelL_mapinO O _ (equiv_pullback_symm k g) _ _). - refine (cancelL_mapinO O _ (g^* k) _ _). - Defined. - - Definition lex_from_ispullback_connmap_mapino_commsq (O : Modality) + (** RSS Theorem 3.1 (vii) implies lex-ness *) + Definition lex_from_ispullback_connmap_mapino_commsq (H : forall {A B C D} (f : A -> B) (g : C -> D) (h : A -> C) (k : B -> D), (IsConnMap O f) -> (IsConnMap O g) -> @@ -214,7 +298,7 @@ Module Lex_Modalities_Theory (Os : Modalities). apply lex_from_isequiv_ismodal_isconnected_types. intros A B f AC BC fM. specialize (H A Unit B Unit (const tt) (const tt) f idmap _ _ _ _ - (fun _ => 1)). + (fun _ => 1)). unfold IsPullback, pullback_corec in H. refine (@isequiv_compose _ _ _ H _ (fun x => x.2.1) _). unfold Pullback. @@ -226,187 +310,58 @@ Module Lex_Modalities_Theory (Os : Modalities). apply (equiv_isequiv (prod_unit_l B)). Defined. - (** 10. Families of modal types indexed by connected types are constant. *) - Definition modal_over_connected_isconst_lex (O : Modality) `{Lex O} - (A : Type) `{IsConnected O A} (P : A -> Type) `{forall x, In O (P x)} - : {Q : Type & In O Q * forall x, P x <~> Q}. - Proof. - exists (O {x:A & P x}); split; [ exact _ | intros x]. - refine (Build_Equiv _ _ (fun p => to O _ (x ; p)) _). - refine (isequiv_conn_map_ino O _). - revert x. - apply conn_map_fiber. - refine (cancelL_conn_map O _ (fun z:{x:A & O {x : A & P x}} => z.2) _ _). - intros z. - refine (isconnected_equiv' O A _ _). - unfold hfiber. - refine (equiv_adjointify (fun x => ((x ; z) ; 1)) - (fun y => y.1.1) _ _). - - intros [[x y] []]; reflexivity. - - intros x; reflexivity. - Defined. - - (** And conversely. *) - Definition lex_from_modal_over_connected_isconst (O : Modality) - (H : forall (A : Type) (P : A -> Type), - (IsConnected O A) -> (forall x, In O (P x)) -> - {Q : Type & In O Q * forall x, P x <~> Q}) - : Lex O. - Proof. - intros A x y ?. - apply isconnected_from_elim_to_O. - (** By assumption, [fun y => O (x = y) : A -> Type_ O] is constant. Thus, [to O (x=x) 1] can be transported around to make it contractible everywhere. *) - specialize (H A (fun z => O (x = z)) _ _). - destruct H as [Q [? H]]. - unfold NullHomotopy. - exists ((H y)^-1 ((H x) (to O _ 1))). - intros []. - symmetry; apply eissect. - Defined. - - (** Lex modalities preserve [n]-types for all [n]. This is definitely not equivalent to lex-ness, since it is true for the truncation modalities that are not lex. But it is also not true of all modalities; e.g. the shape modality in a cohesive topos can take 0-types to [oo]-types. *) - Global Instance istrunc_O_lex `{Funext} {O : Modality} `{Lex O} - {n} {A} `{IsTrunc n A} - : IsTrunc n (O A). - Proof. - generalize dependent A; simple_induction n n IHn; intros A ?. - - exact _. (** Already proven for all modalities. *) - - refine (O_ind (fun x => forall y, IsTrunc n (x = y)) _); intros x. - refine (O_ind (fun y => IsTrunc n (to O A x = y)) _); intros y. - refine (trunc_equiv _ (O_path_cmp O x y)). - Defined. - -End Lex_Modalities_Theory. +End ImpliesLex. (** ** Lex reflective subuniverses *) (** A reflective subuniverse that preserves fibers is in fact a modality (and hence lex). *) -Module Type Preserves_Fibers (Os : ReflectiveSubuniverses). - - Export Os. - Module Export Os_Theory := ReflectiveSubuniverses_Theory Os. - - Parameter isequiv_O_functor_hfiber : - forall (O : ReflectiveSubuniverse) {A B} (f : A -> B) (b : B), - IsEquiv (O_functor_hfiber O f b). - Existing Instance isequiv_O_functor_hfiber. - -End Preserves_Fibers. - -Module Lex_Reflective_Subuniverses - (Os : ReflectiveSubuniverses) (Opf : Preserves_Fibers Os) - <: SigmaClosed Os. - - Import Opf. - - Definition inO_sigma@{u a i j k} (O : ReflectiveSubuniverse@{u a}) - (A:Type@{i}) (B:A -> Type@{j}) - (A_inO : In@{u a i} O A) - (B_inO : forall a, In@{u a j} O (B a)) - : In@{u a k} O {x:A & B x}. - Proof. - pose (g := O_rec@{u a k i k k i} pr1 : O {x : A & B x} -> A). - transparent assert (p : (forall x, g (to O _ x) = x.1)). - { intros x; subst g; apply O_rec_beta. } - apply inO_isequiv_to_O@{u a k k}. - apply isequiv_fcontr; intros x. - refine (contr_equiv' _ (hfiber_hfiber_compose_map@{k k i k k k k k} _ g x)). - apply fcontr_isequiv. - unfold hfiber_compose_map. - transparent assert (h : (Equiv@{k k} (hfiber@{k i} (@pr1 A B) (g x)) - (hfiber@{k i} g (g x)))). - { refine (_ oE equiv_to_O@{u a k k} O _). - refine (_ oE Build_Equiv _ _ - (O_functor_hfiber O (@pr1 A B) (g x)) _). - unfold hfiber. - apply equiv_functor_sigma_id. intros y; cbn. - refine (_ oE (equiv_moveR_equiv_V _ _)). - apply equiv_concat_l. - apply moveL_equiv_V. - unfold g, O_functor. - revert y; apply O_indpaths@{u a k i i k k}; intros [a q]; cbn. - refine (_ @ (O_rec_beta _ _)^). - apply ap, O_rec_beta. } - refine (isequiv_homotopic (h oE equiv_hfiber_homotopic _ _ p (g x)) _). - intros [[a b] q]; cbn. clear h. - unfold O_functor_hfiber. - rewrite O_rec_beta. - unfold functor_sigma; cbn. - refine (path_sigma' _ 1 _). - rewrite O_indpaths_beta; cbn. - unfold moveL_equiv_V, moveR_equiv_V. - Open Scope long_path_scope. - Local Opaque eissect. (* work around bug 4533 *) - set (k := @eissect); change @eissect with k; subst k. (* work around bug 4543 *) - rewrite !ap_pp, !concat_p_pp, !ap_V. - unfold to_O_natural. - rewrite concat_pV_p. - subst p. - rewrite concat_pp_V. - rewrite concat_pp_p; apply moveR_Vp. - rewrite <- !(ap_compose (to O A) (to O A)^-1). - rapply @concat_A1p. - Local Transparent eissect. (* work around bug 4533 *) - Close Scope long_path_scope. - Qed. - -End Lex_Reflective_Subuniverses. - -(** ** Accessible lex modalities *) - -(** We now restrict to lex modalities that are also accessible. *) -Module Accessible_Lex_Modalities_Theory - (Os : Modalities) - (Acc : Accessible_Modalities Os). - - Module Export Acc_Theory := Accessible_Modalities_Theory Os Acc. - Module Export Lex_Theory := Lex_Modalities_Theory Os. - - (** Unfortunately, another subtlety of modules bites us here. It appears that each application of a parametrized module to arguments creates a *new* module, and Coq has no algorithm (not even syntactic identity) for considering two such modules "the same". In particular, the applications [Module Os_Theory := Modalities_Theory Os] that occur in both [Accessible_Modalities_Theory Os Acc] and [Lex_Modalities_Theory Os] create two *different* modules, which appear here as [Acc_Theory.Os_Theory] and [Lex_Theory.Os_Theory]. Thus, for instance, we have two different definitions [Acc_Theory.Os_Theory.O_ind] and [Lex_Theory.Os_Theory.O_ind], etc. - - Fortunately, since these duplicate pairs of definitions each have the same body *and are (usually) transparent*, Coq is willing to consider them identical. Thus, this doesn't cause a great deal of trouble. However, there are certain contexts in which this doesn't apply. For instance, if any definition in [Modalities_Theory] is opaque, then Coq will be unable to notice that its duplicate copies in [Acc_Theory.Os_Theory] and [Lex_Theory.Os_Theory] were identical, potentially causing problems. But since we generally only make definitions opaque if we aren't going to depend on their actual value anywhere else, this is unlikely to be much of an issue. - - A more serious issue is that there are some declarations that function up to a syntactic equality that is stricter than judgmental conversion. For instance, [Inductive] and [Record] definitions, like modules, always create a new object not convertible to any previously existing one. There are no [Inductive] or [Record] definitions in [Modalities_Theory], but there are [Class] declarations, and these function similarly. In particular, typeclass search is unable to use [Instance]s defined in [Acc_Theory] to instantiate typeclasses from [Modalities_Theory] (such as [IsConnected]) needed by functions in [Lex_Theory], and vice versa. - - Fortunately, all the typeclasses defined in [Modalities_Theory] are *singleton* or *definitional* classes (defined with `:= unique_field` rather than `{ field1 ; field2 ; ... }`), which means that they do not actually introduce a new record wrapper. Thus, the [Instance]s from [Acc_Theory] can in fact be typechecked to *belong* to the typeclasses needed by [Lex_Theory], and hence can be supplied (or [assert]ed) explicitly. *) - - (** Probably the most important thing about an accessible lex modality is that the universe of modal types is again modal. Here by "the universe" we mean a universe large enough to contain the generating family; this is why we need accessibility. *) - Global Instance inO_typeO `{Univalence} (O : Modality) `{Lex O} - : In O (Type_ O). - Proof. - apply (snd (inO_iff_isnull O _)); intros i n; simpl in *. - assert (Lex_Theory.Os_Theory.RSU.IsConnected O (acc_gen O i)) - by exact (isconnected_acc_gen O i). - destruct n; [ exact tt | split ]. - - intros P. - (** The case [n=0] is basically just one of the above characterizations of lex-ness. *) - destruct (modal_over_connected_isconst_lex O (acc_gen O i) P) - as [Q [QinO f]]. - exists (fun _ => (Q ; QinO)). - intros x; symmetry; apply path_TypeO. - refine (path_universe (f x)). - - intros A B. - (** The case [n>0] is actually quite easy, using univalence and the fact that modal types are closed under [Equiv]. *) - refine (extendable_postcompose' n _ _ _ - (fun b => (equiv_path_TypeO O (A b) (B b)) - oE (equiv_path_universe (A b) (B b))) - _). - refine (extendable_conn_map_inO O n (@const (acc_gen O i) Unit tt) - (fun b => A b <~> B b)). - (** Typeclass magic! *) - Defined. - - (** [inO_typeO] is also an equivalent characterization of lex-ness for a modality, by the converse of the characterization of lex-ness we used above. *) - Definition lex_inO_typeO (O : Modality) `{In O (Type_ O)} - : Lex O. - Proof. - apply lex_from_modal_over_connected_isconst. - intros A P ? PO. - destruct (isconnected_elim O (Type_ O) (fun x => (P x ; PO x))) - as [Q f]. - exists Q; split; [ exact _ | intros x ]. - apply equiv_path. - exact (ap pr1 (f x)). - Defined. - -End Accessible_Lex_Modalities_Theory. +Definition ismodality_isequiv_O_functor_hfiber (O : ReflectiveSubuniverse) + (H : forall {A B : Type} (f : A -> B) (b : B), + IsEquiv (O_functor_hfiber O f b)) + : IsModality O. +Proof. + intros A'; rapply reflectsD_from_inO_sigma. + intros B B_inO. + pose (A := O A'). + pose (g := O_rec pr1 : O {x : A & B x} -> A). + transparent assert (p : (forall x, g (to O _ x) = x.1)). + { intros x; subst g; apply O_rec_beta. } + apply inO_isequiv_to_O. + apply isequiv_fcontr; intros x. + refine (contr_equiv' _ (hfiber_hfiber_compose_map _ g x)). + apply fcontr_isequiv. + unfold hfiber_compose_map. + transparent assert (h : (hfiber (@pr1 A B) (g x) <~> hfiber g (g x))). + { refine (_ oE equiv_to_O O _). + refine (_ oE Build_Equiv _ _ + (O_functor_hfiber O (@pr1 A B) (g x)) _). + unfold hfiber. + apply equiv_functor_sigma_id. intros y; cbn. + refine (_ oE (equiv_moveR_equiv_V _ _)). + apply equiv_concat_l. + apply moveL_equiv_V. + unfold g, O_functor. + revert y; apply O_indpaths; intros [a q]; cbn. + refine (_ @ (O_rec_beta _ _)^). + apply ap, O_rec_beta. } + refine (isequiv_homotopic (h oE equiv_hfiber_homotopic _ _ p (g x)) _). + intros [[a b] q]; cbn. clear h. + unfold O_functor_hfiber. + rewrite O_rec_beta. + unfold functor_sigma; cbn. + refine (path_sigma' _ 1 _). + rewrite O_indpaths_beta; cbn. + unfold moveL_equiv_V, moveR_equiv_V. + Open Scope long_path_scope. + Local Opaque eissect. (* work around bug 4533 *) + rewrite !ap_pp, !concat_p_pp, !ap_V. + unfold to_O_natural. + rewrite concat_pV_p. + subst p. + rewrite concat_pp_V. + rewrite concat_pp_p; apply moveR_Vp. + rewrite <- !(ap_compose (to O A) (to O A)^-1). + rapply @concat_A1p. + Local Transparent eissect. (* work around bug 4533 *) + Close Scope long_path_scope. +Qed. diff --git a/theories/Modalities/Localization.v b/theories/Modalities/Localization.v index 3c35c71b12f..b74dd0f480d 100644 --- a/theories/Modalities/Localization.v +++ b/theories/Modalities/Localization.v @@ -4,7 +4,6 @@ Require Import HoTT.Basics HoTT.Types HoTT.Cubical. Require Import Extensions EquivalenceVarieties. Require Import ReflectiveSubuniverse Accessible. -Require Import Homotopy.Suspension. Local Open Scope nat_scope. Local Open Scope path_scope. @@ -264,25 +263,18 @@ Definition islocal_equiv_islocal (f : LocalGenerators@{a}) Proof. intros i. (** We have to fiddle with the max universes to get this to work, since [ooextendable_postcompose] requires the max universe in both cases to be the same, whereas we don't want to assume that the hypothesis and conclusion are related in any way. *) - apply lift_ooextendablealong@{a a j k j'}. + apply lift_ooextendablealong@{a a a a a a j j j k j'}. refine (ooextendable_postcompose@{a a i j k k k k k k} _ _ (f i) (fun _ => g) _). - apply lift_ooextendablealong@{a a i i' k}. + apply lift_ooextendablealong@{a a a a a a i i i i' k}. apply Xloc. Defined. -(** A basic operation on local generators is the pointwise suspension. *) -Definition susp_localgen (f : LocalGenerators@{a}) : LocalGenerators@{a}. -Proof. - econstructor; intros i. - exact (functor_susp@{a a a a a a a a a a a a} (f i)). -Defined. - (** ** Localization as a HIT *) Module Export LocalizationHIT. - Private Inductive Localize (f : LocalGenerators@{a}) (X : Type@{i}) + Cumulative Private Inductive Localize (f : LocalGenerators@{a}) (X : Type@{i}) : Type@{max(a,i)} := | loc : X -> Localize f X. @@ -352,96 +344,24 @@ Section Localization. End Localization. -(** We define a wrapper around [LocalGenerators]. See the comments in HIT/Truncations for an explanation. Unlike there, here we make the wrapper a [Record] rather than a [Definition], so that a projection has to be inserted to convert one to the other. This forces us to write [Loc f] as a parameter to all reflective-subuniverse functions, which is really entirely reasonable; we only allowed ourselves to write [n] rather than [Tr n] in the case of truncations because things like connectedness are traditionally defined only for the truncation modality, so users may prefer not to have to think about the fact that there is a modality present. *) -Record Localization_ReflectiveSubuniverse := Loc { unLoc : LocalGenerators }. - -Module Localization_ReflectiveSubuniverses <: ReflectiveSubuniverses. - - (** Here we have to use the extra universe parameter that we built into the definition of [ReflectiveSubuniverse]: a set of localization generators is parametrized by the universe that it lives in, but also by the universe that the *generators* live in, which must be smaller. *) - Definition ReflectiveSubuniverse : Type@{u} := Localization_ReflectiveSubuniverse@{a}. - Check ReflectiveSubuniverse@{u a}. - - Definition O_reflector : ReflectiveSubuniverse@{u a} -> Type@{i} -> Type@{i} - := fun O => Localize@{a i} (unLoc O). - - Definition In : ReflectiveSubuniverse@{u a} -> Type@{i} -> Type@{i} - := fun O X => IsLocal@{i i a} (unLoc O) X. - - Definition O_inO : - forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - In O (O_reflector O T) - := fun O => islocal_localize@{a i i} (unLoc O). - - Definition to : - forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), T -> O_reflector O T - := fun O => @loc@{a i} (unLoc O). - - Definition inO_equiv_inO : - forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}) (U : Type@{j}), - In@{u a i} O T -> forall f : T -> U, IsEquiv f -> - In@{u a j} O U - := fun O => islocal_equiv_islocal@{a i j i j k} (unLoc O). - - Definition hprop_inO@{u a i} `{Funext} - (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}) - : IsHProp (In@{u a i} O T). - Proof. - apply (@trunc_forall@{a i i} _); intros i. - apply ishprop_ooextendable@{a a i i i i i i i i i i i}. - Defined. - - Definition extendable_to_O - (O : ReflectiveSubuniverse@{u a}) {P : Type@{i}} - {Q : Type@{j}} {Q_inO : In@{u a j} O Q} - : ooExtendableAlong@{i i j k} (to O P) (fun _ => Q). - Proof. - apply ext_localize_ind@{a i j i k}; intros ?. +Definition Loc@{a i} (f : LocalGenerators@{a}) : ReflectiveSubuniverse@{i}. +Proof. + simple notypeclasses refine (Build_ReflectiveSubuniverse + (Build_Subuniverse (IsLocal f) _ _) + (fun A => Build_PreReflects _ A (Localize f A) _ (@loc f A)) + (fun A => Build_Reflects _ _ _ _)). + - (** Typeclass inference can find this, but we give it explicitly to prevent extra universes from cropping up. *) + intros ? T; unfold IsLocal. + nrefine (trunc_forall@{a i i}); try assumption. + intros i. + apply ishprop_ooextendable@{a a i i i i i i i i i i i}. + - apply islocal_equiv_islocal. + - apply islocal_localize. + - cbn. intros Q Q_inO. + apply ext_localize_ind; intros ?. apply ooextendable_over_const. apply Q_inO. - Defined. - - (** The separated subuniverse corresponding to localization at [f] is localization at the pointwise suspension of [f]. This (and the following two definitions that prove it) is CORS Lemma 2.15, except that we can't use exactly the same proof since we have to avoid funext. Instead we use the non-funext version proven using wild 0-groupoids in [Homotopy.Suspension]. *) - Definition IsSepFor@{u a} (O' O : ReflectiveSubuniverse@{u a}) : Type@{u} - := paths@{u} (Loc (susp_localgen (unLoc O))) O'. - - Definition inO_paths_from_inSepO@{u a i iplus} - (O' O : ReflectiveSubuniverse@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A) - : In@{u a i} O (x = y). - Proof. - destruct O as [f]; destruct sep; unfold In, IsLocal in *; intros i; cbn in *. - specialize (A_inO i). - (** Unfortunately, [ooextendable_iff_functor_susp] has a ridiculous number of universe parameters. Fortunately, most of them can be [i], and it's fairly easy to figure out which have to be [a]. But a few of them have to be something strictly larger than [i], which is why we included such a universe parameter [iplus] in these fields of [ReflectiveSubuniverse]. *) - assert (e := fst (ooextendable_iff_functor_susp@{a a i i a a a i i iplus i a a a i i iplus i i i i i i i i i i i i iplus i i i iplus i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i} (f i) _) A_inO (x,y)). - cbn in e. - refine (ooextendable_postcompose' _ _ _ _ e). - intros b. - symmetry; apply equiv_dp_const. - Defined. - - Definition inSepO_from_inO_paths@{u a i iplus} - (O' O : ReflectiveSubuniverse@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (e : forall (x y : A), In@{u a i} O (x = y)) - : In@{u a i} O' A. - Proof. - destruct O as [f]; destruct sep; unfold In, IsLocal in *; intros i; cbn in *. - apply (ooextendable_iff_functor_susp@{a a i i a a a i i iplus i a a a i i iplus i i i i i i i i i i i i iplus i i i iplus i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i} (f i)). - intros [x y]. - refine (ooextendable_postcompose' _ _ _ _ (e x y i)). - intros b. - apply equiv_dp_const. - Defined. - -End Localization_ReflectiveSubuniverses. - - -(** If you import the following module [LocM], then you can call all the reflective subuniverse functions with a [LocalGenerators] as the parameter. *) -Module Import LocM := ReflectiveSubuniverses_Theory Localization_ReflectiveSubuniverses. -(** If you don't import it, then you'll need to write [LocM.function_name]. *) -Export LocM.Coercions. - -Coercion Localization_ReflectiveSubuniverse_to_ReflectiveSubuniverse := idmap - : Localization_ReflectiveSubuniverse -> ReflectiveSubuniverse. +Defined. (** Here is the "real" definition of the notation [IsLocal]. Defining it this way allows it to inherit typeclass inference from [In], unlike (for instance) the slightly annoying case of [IsTrunc n] versus [In (Tr n)]. *) Notation IsLocal f := (In (Loc f)). @@ -498,38 +418,97 @@ Arguments local_indpaths_beta : simpl never. (** ** Localization and accessibility *) -(** Localization subuniverses are accessible, essentially by definition. *) -Module Accessible_Localization - <: Accessible_ReflectiveSubuniverses Localization_ReflectiveSubuniverses. - - Import Localization_ReflectiveSubuniverses. - - Definition acc_gen : ReflectiveSubuniverse -> LocalGenerators - := unLoc. - - Definition inO_iff_islocal - (O : ReflectiveSubuniverse@{u a}) (X : Type@{i}) - : iff@{i i i} (In O X) (IsLocal (acc_gen O) X) - := (idmap , idmap). +(** Localization subuniverses are accessible, essentially by definition. Without the universe annotations, [a] and [i] get collapsed. *) +Global Instance accrsu_loc@{a i} (f : LocalGenerators@{a}) : IsAccRSU@{a i} (Loc@{a i} f). +Proof. + unshelve econstructor. + - exact f. + - intros; split; apply idmap. +Defined. -End Accessible_Localization. +(** Conversely, if a subuniverse is accessible, then the corresponding localization subuniverse is equivalent to it, and moreover exists at every universe level and satisfies its computation rules judgmentally. This is called [lift_accrsu] but in fact it works equally well to *lower* the universe level, as long as both levels are no smaller than the size [a] of the generators. *) +Definition lift_accrsu@{a i j} (O : Subuniverse@{i}) `{IsAccRSU@{a i} O} + : ReflectiveSubuniverse@{j} + := Loc@{a j} (acc_lgen O). -(** Conversely, if a reflective subuniverse is accessible, then it can be "nudged" to an equivalent localization. The nudged version has the advantages of satisfying its computation rules judgmentally. *) +(** The lifted universe agrees with the original one, on any universe contained in both [i] and [j] *) +Global Instance O_eq_lift_accrsu@{a i j k} (O : Subuniverse@{i}) `{IsAccRSU@{a i} O} + : O_eq@{i j k} O (lift_accrsu@{a i j} O). +Proof. + (** Anyone stepping through this proof should do [Set Printing Universes]. *) + split; intros A A_inO. + - intros i. + assert (e := fst (inO_iff_islocal O A) A_inO i). + apply (lift_ooextendablealong@{a a a a a a i j k i j} (acc_lgen O i) (fun _ => A)). + exact e. + - apply (inO_iff_islocal O). + intros i. + pose (e := A_inO i). + apply (lift_ooextendablealong@{a a a a a a j i k j i} (acc_lgen O i) (fun _ => A)). + exact e. +Defined. -Module Nudge_ReflectiveSubuniverses - (Os : ReflectiveSubuniverses) - (Acc : Accessible_ReflectiveSubuniverses Os). +Definition O_leq_lift_accrsu@{a i1 i2} + (O1 : ReflectiveSubuniverse@{i1}) (O2 : ReflectiveSubuniverse@{i2}) `{IsAccRSU@{a i1} O1} + `{O_leq@{i1 i2 i2} O1 O2} + : O_leq@{i2 i2 i2} (lift_accrsu@{a i1 i2} O1) O2. +Proof. + intros B B_inO1. + apply (inO_leq@{i1 i2 i2} O1 O2). + apply (snd (inO_iff_islocal O1 B)). + intros i. specialize (B_inO1 i). + apply (lift_ooextendablealong@{a a a a a a i2 i1 i2 i2 i1} (acc_lgen O1 i) (fun _ => B)). + exact B_inO1. +Defined. - (** Once again, we are annoyed that "application of modules is restricted to paths". *) - Module Data <: ReflectiveSubuniverses_Restriction_Data - Localization_ReflectiveSubuniverses. - Definition New_ReflectiveSubuniverse := Os.ReflectiveSubuniverse. - Definition ReflectiveSubuniverses_restriction - : New_ReflectiveSubuniverse -> Localization_ReflectiveSubuniverses.ReflectiveSubuniverse - := fun O => Loc (Acc.acc_gen O). - End Data. +(** Similarly, because localization is a HIT that has an elimination rule into types in *all* universes, for accessible reflective subuniverses we can show that containment implies connectedness properties with the universe containments in the other order. *) +Definition isconnected_O_leq'@{a i1 i2} + (O1 : ReflectiveSubuniverse@{i1}) (O2 : ReflectiveSubuniverse@{i2}) `{IsAccRSU@{a i1} O1} + (** Compared to [O_leq@{i1 i2 i1}] and [A : Type@{i1}] in [isconnected_O_leq], these two lines are what make [i2 <= i1] instead of vice versa. *) + `{O_leq@{i1 i2 i2} O1 O2} (A : Type@{i2}) + `{IsConnected O2 A} + : IsConnected O1 A. +Proof. + (** Anyone stepping through this proof should do [Set Printing Universes]. *) + srefine (isconnected_O_leq O1 (lift_accrsu@{a i1 i1} O1) A). + 1-2:exact _. + change (Contr@{i1} (Localize@{a i2} (acc_lgen@{a i1} O1) A)). + (** At this point you should also do [Unset Printing Notations] to see the universe annotation on [IsTrunc] change. *) + refine (contr_equiv'@{i2 i1} _ 1%equiv). + change (IsConnected@{i2} (lift_accrsu@{a i1 i2} O1) A). + srapply (isconnected_O_leq _ O2). + rapply O_leq_lift_accrsu. +Defined. - Module Nudged <: ReflectiveSubuniverses - := ReflectiveSubuniverses_Restriction Localization_ReflectiveSubuniverses Data. +(** And similarly for connected maps. *) +Definition conn_map_O_leq'@{a i1 i2} + (O1 : ReflectiveSubuniverse@{i1}) (O2 : ReflectiveSubuniverse@{i2}) `{IsAccRSU@{a i1} O1} + `{O_leq@{i1 i2 i2} O1 O2} {A B : Type@{i2}} + (f : A -> B) `{IsConnMap O2 A B f} + : IsConnMap O1 f. +Proof. + (** Anyone stepping through this proof should do [Set Printing Universes]. *) + intros b. + apply (isconnected_equiv' O1 (hfiber@{i2 i2} f b)). + - srapply equiv_adjointify. + 1-2:intros [u p]; exact (u;p). + all:intros [u p]; reflexivity. + - apply (isconnected_O_leq' O1 O2). + apply isconnected_hfiber_conn_map. +Defined. -End Nudge_ReflectiveSubuniverses. +(** The same is true for inverted maps, too. *) +Definition O_inverts_O_leq'@{a i1 i2} + (O1 : ReflectiveSubuniverse@{i1}) (O2 : ReflectiveSubuniverse@{i2}) `{IsAccRSU@{a i1} O1} + `{O_leq@{i1 i2 i2} O1 O2} {A B : Type@{i2}} + (f : A -> B) `{O_inverts O2 f} + : O_inverts O1 f. +Proof. + assert (oleq := O_leq_lift_accrsu O1 O2). + assert (e := O_inverts_O_leq (lift_accrsu@{a i1 i2} O1) O2 f); clear oleq. + nrapply (O_inverts_O_leq O1 (lift_accrsu@{a i1 i1} O1) f). + 1:exact _. + (** It looks like we can say [exact e], but that would collapse the universes [i1] and [i2]. You can check with [Set Printing Universes. Unset Printing Notations.] that [e] and the goal have different universes. So instead we do this: *) + refine (@isequiv_homotopic _ _ _ _ e _). + apply O_indpaths; intros x; reflexivity. +Defined. diff --git a/theories/Modalities/Modality.v b/theories/Modalities/Modality.v index 903e0130eae..e8cbed22b8e 100644 --- a/theories/Modalities/Modality.v +++ b/theories/Modalities/Modality.v @@ -8,509 +8,470 @@ Local Open Scope path_scope. (** * Modalities *) -Module Type Modalities. - - Parameter Modality@{u a} : Type2@{u a}. +(** ** Dependent eliminators *) + +(** A dependent version of the reflection universal property. For later use we generalize it to refer to different subuniverses in the reflection and the elimination target. *) +Class ReflectsD@{i} (O' O : Subuniverse@{i}) (T : Type@{i}) + `{PreReflects@{i} O' T} := +{ + extendable_to_OO : + forall (Q : O_reflector O' T -> Type@{i}) {Q_inO : forall x, In O (Q x)}, + ooExtendableAlong (to O' T) Q +}. + +(** In particular, from this we get a dependent eliminator. *) +Definition OO_ind {O' : Subuniverse} (O : Subuniverse) + {A : Type} `{ReflectsD O' O A} + (B : O_reflector O' A -> Type) {B_inO : forall oa, In O (B oa)} + (f : forall a, B (to O' A a)) (oa : O_reflector O' A) + : B oa + := (fst (extendable_to_OO B 1%nat) f).1 oa. + +Definition OO_ind_beta {O' O : Subuniverse} {A : Type} `{ReflectsD O' O A} + (B : O_reflector O' A -> Type) {B_inO : forall oa, In O (B oa)} + (f : forall a, B (to O' A a)) (a : A) + : OO_ind O B f (to O' A a) = f a + := (fst (extendable_to_OO B 1%nat) f).2 a. + +(** Conversely, if [O] is closed under path-types, a dependent eliminator suffices to prove the whole dependent universal property. *) +Definition reflectsD_from_OO_ind@{i} {O' O : Subuniverse@{i}} + {A : Type@{i}} `{PreReflects O' A} + (OO_ind' : forall (B : O_reflector O' A -> Type@{i}) + (B_inO : forall oa, In O (B oa)) + (f : forall a, B (to O' A a)) + oa, B oa) + (OO_ind_beta' : forall (B : O_reflector O' A -> Type@{i}) + (B_inO : forall oa, In O (B oa)) + (f : forall a, B (to O' A a)) + a, OO_ind' B B_inO f (to O' A a) = f a) + (inO_paths' : forall (B : Type@{i}) (B_inO : In O B) + (z z' : B), In O (z = z')) + : ReflectsD O' O A. +Proof. + constructor. + intros Q Q_inO n. + revert Q Q_inO. simple_induction n n IHn; intros Q Q_inO. + 1:exact tt. + split. + - intros g. + exists (OO_ind' Q _ g). + rapply OO_ind_beta'. + - intros h k. + rapply IHn. +Defined. - (** These are the same as for a reflective subuniverse. *) +(** In particular, this is the case if [O] is a reflective subuniverse. *) +Definition reflectsD_from_RSU {O' : Subuniverse} {O : ReflectiveSubuniverse} + {A : Type} `{PreReflects O' A} + (OO_ind' : forall (B : O_reflector O' A -> Type) + (B_inO : forall oa, In O (B oa)) + (f : forall a, B (to O' A a)) + oa, B oa) + (OO_ind_beta' : forall (B : O_reflector O' A -> Type) + (B_inO : forall oa, In O (B oa)) + (f : forall a, B (to O' A a)) + a, OO_ind' B B_inO f (to O' A a) = f a) + : ReflectsD O' O A + := reflectsD_from_OO_ind OO_ind' OO_ind_beta' _. + +(** Of course, with funext this becomes an actual equivalence. *) +Definition isequiv_oD_to_O + {fs : Funext} (O' O : Subuniverse) {A : Type} `{ReflectsD O' O A} + (B : O_reflector O' A -> Type) `{forall a, In O (B a)} + : IsEquiv (fun (h : forall oa, B oa) => h oD to O' A). +Proof. + apply isequiv_ooextendable, extendable_to_OO; assumption. +Defined. - Parameter O_reflector@{u a i} : forall (O : Modality@{u a}), - Type2le@{i a} -> Type2le@{i a}. - Parameter In@{u a i} : forall (O : Modality@{u a}), - Type2le@{i a} -> Type2le@{i a}. +(** ** The strong order *) - Parameter O_inO@{u a i} : forall (O : Modality@{u a}) (T : Type@{i}), - In@{u a i} O (O_reflector@{u a i} O T). +(** Note the reversal of the order: [O1 << O2] means that [O2] has dependent eliminators into [O1]. *) +Class O_strong_leq (O1 O2 : ReflectiveSubuniverse) + := reflectsD_strong_leq : forall A, ReflectsD O2 O1 A. +Global Existing Instance reflectsD_strong_leq. - Parameter to@{u a i} : forall (O : Modality@{u a}) (T : Type@{i}), - T -> O_reflector@{u a i} O T. +Notation "O1 << O2" := (O_strong_leq O1 O2) (at level 70) : subuniverse_scope. +Open Scope subuniverse_scope. - Parameter inO_equiv_inO@{u a i j k} : - forall (O : Modality@{u a}) (T : Type@{i}) (U : Type@{j}) - (T_inO : In@{u a i} O T) (f : T -> U) (feq : IsEquiv f), - (** We add an extra universe parameter that's bigger than both [i] and [j]. This seems to be necessary for the proof of repleteness in some examples, such as easy modalities. *) - let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - In@{u a j} O U. +(** The strong order implies the weak order. *) +Global Instance O_leq_strong_leq {O1 O2 : ReflectiveSubuniverse} `{O1 << O2} + : O1 <= O2. +Proof. + intros A A_inO1. + srapply inO_to_O_retract. + - exact (OO_ind O1 (fun _:O2 A => A) idmap). + - intros a. srapply OO_ind_beta. +Defined. - Parameter hprop_inO@{u a i} - : Funext -> forall (O : Modality@{u a}) (T : Type@{i}), - IsHProp (In@{u a i} O T). +(** The strong order is not obviously transitive, but it composes with the weak order on one side at least. *) +Definition O_strong_leq_trans_l (O1 O2 O3 : ReflectiveSubuniverse) + `{O1 <= O2} `{O2 << O3} + : O1 << O3. +Proof. + intros A; constructor; intros B B_inO. + apply (extendable_to_OO (O := O2)). + intros x. + srapply inO_leq; apply B_inO. +Defined. - Parameter IsSepFor@{u a} : forall (O' O : Modality@{u a}), Type@{u}. - Parameter inO_paths_from_inSepO@{u a i iplus} - : forall (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A), - let gt := (Type1@{i} : Type@{iplus}) in - In@{u a i} O (x = y). +(** ** Modalities *) - Parameter inSepO_from_inO_paths@{u a i iplus} - : forall (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}), - let gt := (Type1@{i} : Type@{iplus}) in - (forall (x y : A), In@{u a i} O (x = y)) -> In@{u a i} O' A. +(** A modality is a reflective subuniverse with a dependent universal property with respect to itself. *) +Notation IsModality O := (O << O). - (** Now instead of [extendable_to_O], we have an ordinary induction principle. *) +(** However, it's not clear what the best bundled definition of modality is. The obvious one [{ O : ReflectiveSubuniverse & IsModality O}] has the advantage that bundling a reflective subuniverse into a modality and then unbundling it is definitionally the identity; but it is redundant, since the dependent universal property implies the non-dependent one, and in practice most modalities are constructed directly with a dependent eliminator. Thus, for now at least, we take the following definition, which in RSS is called a "uniquely eliminating modality". *) +Record Modality@{i} := Build_Modality' +{ + modality_subuniv : Subuniverse@{i} ; + modality_prereflects : forall (T : Type@{i}), + PreReflects modality_subuniv T ; + modality_reflectsD : forall (T : Type@{i}), + ReflectsD modality_subuniv modality_subuniv T ; +}. - (** Unfortunately, we have to define these parameters as [_internal] versions and redefine them later. This is because eventually we want the [In] fields of [O_ind] and [O_ind_beta] to refer to the [In] typeclass defined for [ReflectiveSubuniverse]s, but in order to prove that modalities *are* reflective subuniverses we need to already have [O_ind] and [O_ind_beta]. *) - Parameter O_ind_internal@{u a i j k} - : forall (O : Modality@{u a}) - (A : Type2le@{i a}) - (B : O_reflector@{u a i} O A -> Type2le@{j a}) - (B_inO : forall oa, In@{u a j} O (B oa)), - (** We add an extra unused universe parameter [k] that's [>= max(i,j)]. This seems to be necessary for some examples, such as [Nullification], which are constructed by way of an operation that requires such a universe. *) - let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - (forall a, B (to O A a)) -> forall a, B a. +Global Existing Instance modality_reflectsD. +(** We don't declare [modality_subuniv] as a coercion or [modality_prereflects] as a global instance, because we want them only to be found by way of the following "unbundling" coercion to reflective subuniverses. *) - Parameter O_ind_beta_internal@{u a i j k} - : forall (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - (f : forall a : A, B (to O A a)) (a:A), - O_ind_internal@{u a i j k} O A B B_inO f (to O A a) = f a. +Definition modality_to_reflective_subuniverse (O : Modality@{i}) + : ReflectiveSubuniverse@{i}. +Proof. + refine (Build_ReflectiveSubuniverse + (modality_subuniv O) (modality_prereflects O) _). + intros T; constructor. + intros Q Q_inO. + srapply extendable_to_OO. +Defined. - Parameter minO_paths@{u a i} - : forall (O : Modality@{u a}) - (A : Type2le@{i a}) (A_inO : In@{u a i} O A) (z z' : A), - In@{u a i} O (z = z'). +Coercion modality_to_reflective_subuniverse : Modality >-> ReflectiveSubuniverse. -End Modalities. +(** Unfortunately, sometimes [modality_subuniv] pops up anyway. The following hint helps typeclass inference look through it. *) +Hint Extern 0 (In (modality_subuniv _) _) => progress change modality_subuniv with (rsu_subuniv o modality_to_reflective_subuniverse) in * : typeclass_instances. -(** ** Modalities are reflective subuniverses *) +(** Modalities are precisely the reflective subuniverses that are [<<] themselves. *) +Global Instance ismodality_modality (O : Modality) : IsModality O. +Proof. + intros A; exact _. +Defined. -(** We show that modalities have underlying reflective subuniverses. It is important in some applications, such as [Trunc], that this construction uses the general [O_ind] given as part of the modality data. For instance, this ensures that [O_functor] reduces to simply an application of [O_ind]. +Definition modality_ismodality (O : ReflectiveSubuniverse) `{IsModality O} : Modality. +Proof. + rapply Build_Modality'. +Defined. - Note also that our choice of how to define reflective subuniverses differently from the book, using [ooExtendableAlong] enables us to prove this without using funext. *) +(** Of course, modalities have dependent eliminators. *) +Definition O_ind {O : Subuniverse} {A : Type} `{ReflectsD O O A} + := @OO_ind O O A _ _. +Arguments O_ind {O A _ _} B {B_inO} f oa. +Definition O_ind_beta {O : Subuniverse} {A : Type} `{ReflectsD O O A} + := @OO_ind_beta O O A _ _. +Arguments O_ind_beta {O A _ _} B {B_inO} f a. + +(** Conversely, as remarked above, we can build a modality from a dependent eliminator as long as we assume the modal types are closed under paths. This is probably the most common way to define a modality, and one might argue that this would be a better definition of the bundled type [Modality]. For now we simply respect that by dignifying it with the unprimed constructor name [Build_Modality]. *) +Definition Build_Modality + (In' : Type -> Type) + (hprop_inO' : Funext -> forall T : Type, IsHProp (In' T)) + (inO_equiv_inO' : forall T U : Type, + In' T -> forall f : T -> U, IsEquiv f -> In' U) + (O_reflector' : Type -> Type) + (O_inO' : forall T, In' (O_reflector' T)) + (to' : forall T, T -> O_reflector' T) + (O_ind' : forall (A : Type) (B : O_reflector' A -> Type) + (B_inO : forall oa, In' (B oa)) + (f : forall a, B (to' A a)) + (z : O_reflector' A), + B z) + (O_ind_beta' : forall (A : Type) (B : O_reflector' A -> Type) + (B_inO : forall oa, In' (B oa)) + (f : forall a, B (to' A a)) (a : A), + O_ind' A B B_inO f (to' A a) = f a) + (inO_paths' : forall (A : Type) (A_inO : In' A) (z z' : A), + In' (z = z')) + : Modality. +Proof. + pose (O := Build_Subuniverse In' hprop_inO' inO_equiv_inO'). + simple refine (Build_Modality' O _ _); intros T. + - exact (Build_PreReflects O T (O_reflector' T) (O_inO' T) (to' T)). + - srapply reflectsD_from_OO_ind. + + rapply O_ind'. + + rapply O_ind_beta'. + + rapply inO_paths'. +Defined. -Module Modalities_to_ReflectiveSubuniverses - (Os : Modalities) <: ReflectiveSubuniverses. +(** When combined with [isequiv_oD_to_O], this yields Theorem 7.7.7 in the book. *) +Definition isequiv_oD_to_O_modality + `{Funext} (O : Modality) {A : Type} + (B : O A -> Type) `{forall a, In O (B a)} + : IsEquiv (fun (h : forall oa, B oa) => h oD to O A). +Proof. + srapply (isequiv_oD_to_O O O). +Defined. - Import Os. - Fixpoint O_extendable@{u a i j k} (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector O A -> Type@{j}) - (B_inO : forall a, In@{u a j} O (B a)) (n : nat) - : ExtendableAlong@{i i j k} n (to O A) B. - Proof. - destruct n as [|n]. - - exact tt. - - split. - + intros g. - exists (O_ind_internal@{u a i j k} O A B B_inO g); intros x. - apply O_ind_beta_internal. - + intros h k. - apply O_extendable; intros x. - apply minO_paths; trivial. - Defined. +(** ** Dependent sums *) - Definition ReflectiveSubuniverse := Modality. - - Definition O_reflector@{u a i} := O_reflector@{u a i}. - (** Work around https://coq.inria.fr/bugs/show_bug.cgi?id=3807 *) - Definition In@{u a i} : forall (O : ReflectiveSubuniverse@{u a}), - Type2le@{i a} -> Type2le@{i a} - := In@{u a i}. - Definition O_inO@{u a i} : forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - In@{u a i} O (O_reflector@{u a i} O T) - := O_inO@{u a i}. - Definition to@{u a i} := to@{u a i}. - Definition inO_equiv_inO@{u a i j k} : - forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}) (U : Type@{j}) - (T_inO : In@{u a i} O T) (f : T -> U) (feq : IsEquiv f), - In@{u a j} O U - := inO_equiv_inO@{u a i j k}. - Definition hprop_inO@{u a i} - : Funext -> forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - IsHProp (In@{u a i} O T) - := hprop_inO@{u a i}. - Definition IsSepFor@{u a} := IsSepFor@{u a}. - Definition inO_paths_from_inSepO@{u a i iplus} := inO_paths_from_inSepO@{u a i iplus}. - Definition inSepO_from_inO_paths@{u a i iplus} := inSepO_from_inO_paths@{u a i iplus}. - - (** Corollary 7.7.8, part 1 *) - Definition extendable_to_O@{u a i j k} (O : ReflectiveSubuniverse@{u a}) - {P : Type2le@{i a}} {Q : Type2le@{j a}} {Q_inO : In@{u a j} O Q} - : ooExtendableAlong@{i i j k} (to O P) (fun _ => Q) - := fun n => O_extendable O P (fun _ => Q) (fun _ => Q_inO) n. - -End Modalities_to_ReflectiveSubuniverses. - - -(** Conversely, if a reflective subuniverse is closed under sigmas, it is a modality. This is a bit annoying to state using modules, but it is not really a problem in practice: in most or all examples, constructing [O_ind] directly is just as easy, and preferable because it sometimes gives a judgmental computation rule. *) - -Module Type SigmaClosed (Os : ReflectiveSubuniverses). - - Import Os. - - Parameter inO_sigma@{u a i j k} - : forall (O : ReflectiveSubuniverse@{u a}) - (A:Type@{i}) (B:A -> Type@{j}) - (A_inO : In@{u a i} O A) - (B_inO : forall a, In@{u a j} O (B a)), - In@{u a k} O {x:A & B x}. - -End SigmaClosed. - -Module ReflectiveSubuniverses_to_Modalities - (Os : ReflectiveSubuniverses) (OsSigma : SigmaClosed Os) - <: Modalities. - - Import Os OsSigma. - Module Import Os_Theory := ReflectiveSubuniverses_Theory Os. - - Definition Modality := ReflectiveSubuniverse. - - Definition O_reflector@{u a i} := O_reflector@{u a i}. - (** Work around https://coq.inria.fr/bugs/show_bug.cgi?id=3807 *) - Definition In@{u a i} := In@{u a i}. - Definition O_inO@{u a i} := @O_inO@{u a i}. - Definition to@{u a i} := to@{u a i}. - Definition inO_equiv_inO@{u a i j k} := @inO_equiv_inO@{u a i j k}. - Definition hprop_inO@{u a i} := hprop_inO@{u a i}. - Definition IsSepFor@{u a} := IsSepFor@{u a}. - Definition inO_paths_from_inSepO@{u a i iplus} := inO_paths_from_inSepO@{u a i iplus}. - Definition inSepO_from_inO_paths@{u a i iplus} := inSepO_from_inO_paths@{u a i iplus}. - - Definition O_ind_internal@{u a i j k} (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector@{u a i} O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - : (forall a, B (to O A a)) -> forall a, B a - := fun g => pr1 ((O_ind_from_inO_sigma@{u a i j k} O (inO_sigma O)) - A B B_inO g). - - Definition O_ind_beta_internal@{u a i j k} (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector@{u a i} O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - (f : forall a : A, B (to O A a)) (a:A) - : O_ind_internal O A B B_inO f (to O A a) = f a - := pr2 ((O_ind_from_inO_sigma@{u a i j k} O (inO_sigma O)) - A B B_inO f) a. - - Definition minO_paths@{u a i} (O : Modality@{u a}) - (A : Type@{i}) (A_inO : In@{u a i} O A) (z z' : A) - : In O (z = z') - := @inO_paths@{u a i i} O A A_inO z z'. - -End ReflectiveSubuniverses_to_Modalities. +(** A dependent elimination of a reflective subuniverse [O'] into [O] implies that the sum of a family of [O]-modal types over an [O']-modal type is [O']-modal. More specifically, for a particular such sum it suffices for the [O']-reflection of that sum to dependently eliminate into [O]. *) +Global Instance inO_sigma_reflectsD + {O' : ReflectiveSubuniverse} {O : Subuniverse} + {A : Type} (B : A -> Type) `{!ReflectsD O' O (sig B)} + `{In O' A} `{forall a, In O (B a)} + : In O' {x:A & B x}. +Proof. + pose (h := fun x => @O_rec O' ({x:A & B x}) A _ _ _ pr1 x). + assert (p := (fun z => O_rec_beta pr1 z) : h o (to O' _) == pr1). + pose (g := fun z => (transport B ((p z)^) z.2)). + simpl in *. + pose (f := OO_ind O (fun x:O' (sig B) => B (h x)) g). + pose (q := OO_ind_beta (fun x:O' (sig B) => B (h x)) g). + apply inO_to_O_retract with (mu := fun w => (h w; f w)). + intros [x1 x2]. + simple refine (path_sigma B _ _ _ _); simpl. + - apply p. + - refine (ap _ (q (x1;x2)) @ _). + unfold g; simpl. exact (transport_pV B _ _). +Defined. +(** Specialized to a modality, this yields the implication (ii) => (i) from Theorem 7.7.4 of the book, and also Corollary 7.7.8, part 2. *) +Global Instance inO_sigma (O : Modality) + {A:Type} (B : A -> Type) `{In O A} `{forall a, In O (B a)} + : In O {x:A & B x} + := _. -(** ** Easy modalities *) +(** This implies that the composite of modal maps is modal. *) +Global Instance mapinO_compose {O : Modality} {A B C : Type} (f : A -> B) (g : B -> C) + `{MapIn O _ _ f} `{MapIn O _ _ g} + : MapIn O (g o f). +Proof. + intros c. + refine (inO_equiv_inO' _ (hfiber_compose f g c)^-1). +Defined. -(** Our definition of modality is slightly different from the one in the book, which requires an induction principle only into families of the form [fun oa => O (B oa)], and similarly only that path-spaces of types [O A] are modal, where "modal" means that the unit is an equivalence. This is equivalent, roughly since every modal type [A] (in this sense) is equivalent to [O A]. +(** It also implies Corollary 7.3.10 from the book, generalized to modalities. (Theorem 7.3.9 is true for any reflective subuniverse; we called it [equiv_O_sigma_O].) *) +Corollary equiv_sigma_inO_O {O : Modality} {A : Type} `{In O A} (P : A -> Type) + : {x:A & O (P x)} <~> O {x:A & P x}. +Proof. + transitivity (O {x:A & O (P x)}). + - rapply equiv_to_O. + - apply equiv_O_sigma_O. +Defined. -However, our definition is more convenient in formalized applications because in some examples (such as [Trunc] and closed modalities), there is a naturally occurring [O_ind] into all modal types that is not judgmentally equal to the one that can be constructed by passing through [O] and back again. Thus, when we apply general theorems about modalities to a particular modality such as [Trunc], the proofs will reduce definitionally to "the way we would have proved them directly" if we didn't know about general modalities. +(** Conversely, if the sum of a particular family of [O]-modal types over an [O']-reflection is in [O'], then that family admits a dependent eliminator. *) +Definition extension_from_inO_sigma + {O' : Subuniverse} (O : Subuniverse) + {A:Type} `{Reflects O' A} (B: O_reflector O' A -> Type) + {inO_sigma : In O' {z:O_reflector O' A & B z}} + (g : forall x, B (to O' A x)) + : ExtensionAlong (to O' A) B g. +Proof. + set (Z := sigT B) in *. + pose (g' := (fun a:A => (to O' A a ; g a)) : A -> Z). + pose (f' := O_rec (O := O') g'). + pose (eqf := (O_rec_beta g') : f' o to O' A == g'). + pose (eqid := O_indpaths (pr1 o f') idmap (fun x => ap@{k i} pr1 (eqf x))). + exists (fun z => transport B (eqid z) ((f' z).2)). + intros a. unfold eqid. + refine (_ @ pr2_path (O_rec_beta g' a)). + refine (ap (fun p => transport B p (O_rec g' (to O' A a)).2) _). + srapply O_indpaths_beta. +Defined. -On the other hand, in other examples (such as [~~] and open modalities) it is easier to construct the latter weaker induction principle. Thus, we now show how to get from that to our definition of modality. *) +(** And even a full equivalence of spaces of sections. This is stated in CORS Proposition 2.8 (but our version avoids funext by using [ooExtendableAlong], as usual). *) +Definition ooextendable_from_inO_sigma + {O' : ReflectiveSubuniverse} (O : Subuniverse) + {A : Type} (B: O_reflector O' A -> Type) + {inO_sigma : In O' {z:O_reflector O' A & B z}} + : ooExtendableAlong (to O' A) B. +Proof. + intros n; generalize dependent A. + induction n as [|n IHn]; intros; [ exact tt | cbn ]. + refine (extension_from_inO_sigma O B , _). + intros h k; nrapply IHn. + set (Z := sigT B) in *. + pose (W := sigT (fun a => B a * B a)). + nrefine (inO_equiv_inO' (Pullback (A := W) (fun a:O_reflector O' A => (a;(h a,k a))) + (fun z:Z => (z.1;(z.2,z.2)))) _). + { refine (inO_pullback O' _ _). + exact (inO_equiv_inO' _ (equiv_sigprod_pullback B B)^-1). } + unfold Pullback. + (** The rest is just contracting a couple of based path spaces. It seems like it should be less work than this. *) + apply equiv_functor_sigma_id; intros z; cbn. + refine (_ oE equiv_functor_sigma_id _). + 2:intros; symmetry; apply equiv_path_sigma. + refine (_ oE (equiv_sigma_assoc _ _)^-1%equiv); cbn. + refine (_ oE equiv_functor_sigma_id _). + 2:intros; apply equiv_sigma_symm. + refine (_ oE (equiv_sigma_assoc' _ _)). + refine (_ oE equiv_contr_sigma _); cbn. + refine (_ oE equiv_functor_sigma_id _). + 2:{ intros; symmetry; etransitivity; revgoals. + - apply equiv_path_prod. + - apply equiv_sigma_prod0. } + refine (_ oE equiv_sigma_assoc' _ _). + refine (_ oE equiv_contr_sigma _); cbn. + apply equiv_path_inverse. +Defined. -Module Type EasyModalities. +(** Thus, if this holds for all sigma-types, we get the dependent universal property. Making this an [Instance] causes typeclass search to spin. Note the slightly different hypotheses, which mean that we can't just use the previous result: here we need only assume that the [O']-reflection of [A] exists rather than that [O'] is fully reflective, at the cost of assuming that [O] is fully reflective (although actually, closed under path-spaces would suffice). *) +Definition reflectsD_from_inO_sigma + {O' : Subuniverse} (O : ReflectiveSubuniverse) + {A : Type} `{Reflects O' A} + (inO_sigma : forall (B: O_reflector O' A -> Type), + (forall oa, In O (B oa)) -> In O' {z:O_reflector O' A & B z}) + : ReflectsD O' O A. +Proof. + constructor; intros B B_inO. + intros n; generalize dependent A. + induction n as [|n IHn]; intros; [ exact tt | cbn ]. + refine (extension_from_inO_sigma O B , _). + intros h k; rapply IHn. +Defined. - Parameter Modality@{u a} : Type2@{u a}. +(** In particular, we get the converse implication (i) => (ii) from Theorem 7.7.4 of the book: a reflective subuniverse closed under sigmas is a modality. *) +Definition modality_from_inO_sigma (O : ReflectiveSubuniverse) + (H : forall (A:Type) (B:A -> Type) + {A_inO : In O A} `{forall a, In O (B a)}, + (In O {x:A & B x})) + : Modality. +Proof. + refine (Build_Modality' O _ _). + intros; srapply reflectsD_from_inO_sigma. +Defined. - Parameter O_reflector@{u a i} : forall (O : Modality@{u a}), - Type2le@{i a} -> Type2le@{i a}. - Parameter to@{u a i} : forall (O : Modality@{u a}) (T : Type@{i}), - T -> O_reflector@{u a i} O T. +(** ** Connectedness of the units *) - Parameter O_indO@{u a i j} - : forall (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}), - (forall a, O_reflector@{u a j} O (B (to O A a))) - -> forall z, O_reflector@{u a j} O (B z). +(** Dependent reflection can also be characterized by connectedness of the unit maps. *) +Global Instance conn_map_to_O_reflectsD {O' : Subuniverse} (O : ReflectiveSubuniverse) + {A : Type} `{ReflectsD O' O A} + : IsConnMap O (to O' A). +Proof. + apply conn_map_from_extension_elim. + intros P P_inO f. + exact (fst (extendable_to_OO (O := O) P 1%nat) f). +Defined. - Parameter O_indO_beta@{u a i j} - : forall (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}) - (f : forall a, O_reflector@{u a j} O (B (to O A a))) a, - O_indO O A B f (to O A a) = f a. +Definition reflectsD_from_conn_map_to_O {O' : Subuniverse} (O : ReflectiveSubuniverse) + {A : Type} `{PreReflects O' A} `{IsConnMap O _ _ (to O' A)} + : ReflectsD O' O A. +Proof. + constructor; rapply ooextendable_conn_map_inO. +Defined. - Parameter minO_pathsO@{u a i} - : forall (O : Modality@{u a}) (A : Type@{i}) - (z z' : O_reflector@{u a i} O A), - IsEquiv (to@{u a i} O (z = z')). +(** In particular, if [O1 << O2] then every [O2]-unit is [O1]-connected. *) +Global Instance conn_map_to_O_strong_leq + {O1 O2 : ReflectiveSubuniverse} `{O1 << O2} (A : Type) + : IsConnMap O1 (to O2 A) + := _. -End EasyModalities. +(** Thus, if [O] is a modality, then every [O]-unit is [O]-connected. This is Corollary 7.5.8 in the book. *) +Global Instance conn_map_to_O {O : Modality} (A : Type) + : IsConnMap O (to O A) + := _. -Module EasyModalities_to_Modalities (Os : EasyModalities) -<: Modalities. - Import Os. +(** ** Easy modalities *) - Definition Modality := Modality. - (** Work around https://coq.inria.fr/bugs/show_bug.cgi?id=3807 *) - Definition O_reflector@{u a i} := O_reflector@{u a i}. - Definition to@{u a i} := to@{u a i}. +(** The book uses yet a different definition of modality, which requires an induction principle only into families of the form [fun oa => O (B oa)], and similarly only that path-spaces of types [O A] are "modal" in the sense that the unit is an equivalence. As shown in section 1 of RSS, this is equivalent, roughly since every modal type [A] (in this sense) is equivalent to [O A]. - Definition In@{u a i} - : forall (O : Modality@{u a}), Type@{i} -> Type@{i} - := fun O A => IsEquiv@{i i} (to O A). +Our definitions are more convenient in formalized applications because in some examples (such as [Trunc] and closed modalities), there is a naturally occurring [O_ind] into all modal types that is not judgmentally equal to the one that can be constructed by passing through [O] and back again. Thus, when we apply general theorems about modalities to a particular modality such as [Trunc], the proofs will reduce definitionally to "the way we would have proved them directly" if we didn't know about general modalities. - Definition hprop_inO@{u a i} `{Funext} (O : Modality@{u a}) - (T : Type@{i}) - : IsHProp (In@{u a i} O T). - Proof. - unfold In. - exact (hprop_isequiv (to O T)). - Defined. +On the other hand, in other examples (such as [~~] and open modalities) it is easier to construct the latter weaker induction principle. Thus, we now show how to get from that to our definition of modality. *) - Definition O_ind_internal@{u a i j k} (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector@{u a i} O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - : let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - (forall a, B (to O A a)) -> forall oa, B oa. +Section EasyModalities. + + Context (O_reflector : Type@{i} -> Type@{i}) + (to : forall (T : Type@{i}), T -> O_reflector T) + (O_indO + : forall (A : Type@{i}) + (B : O_reflector A -> Type@{i}) + (f : forall a, O_reflector (B (to A a))) + (z : O_reflector A), + O_reflector (B z)) + (O_indO_beta + : forall (A : Type@{i}) + (B : O_reflector A -> Type@{i}) + (f : forall a, O_reflector (B (to A a))) (a : A), + O_indO A B f (to A a) = f a) + (inO_pathsO + : forall (A : Type@{i}) (z z' : O_reflector A), + IsEquiv (to (z = z'))). + + Local Definition In_easy : Type@{i} -> Type@{i} + := fun A => IsEquiv (to A). + + Local Definition O_ind_easy + (A : Type) (B : O_reflector A -> Type) + (B_inO : forall oa, In_easy (B oa)) + : (forall a, B (to A a)) -> forall oa, B oa. Proof. simpl; intros f oa. - pose (H := B_inO oa); unfold In in H. - apply ((to O (B oa))^-1). + pose (H := B_inO oa); unfold In_easy in H. + apply ((to (B oa))^-1). apply O_indO. intros a; apply to, f. Defined. - Definition O_ind_beta_internal@{u a i j k} (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector@{u a i} O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - (f : forall a : A, B (to O A a)) (a:A) - : O_ind_internal@{u a i j k} O A B B_inO f (to O A a) = f a. + Local Definition O_ind_easy_beta + (A : Type) (B : O_reflector A -> Type) + (B_inO : forall oa, In_easy (B oa)) + (f : forall a : A, B (to A a)) (a:A) + : O_ind_easy A B B_inO f (to A a) = f a. Proof. - unfold O_ind_internal. + unfold O_ind_easy. apply moveR_equiv_V. - apply @O_indO_beta with (f := fun x => to O _ (f x)). + apply @O_indO_beta with (f := fun x => to _ (f x)). Qed. - Definition O_inO@{u a i} (O : Modality@{u a}) (A : Type@{i}) - : In@{u a i} O (O_reflector@{u a i} O A). + Local Definition O_inO_easy (A : Type) : In_easy (O_reflector A). Proof. - refine (isequiv_adjointify (to O (O_reflector O A)) - (O_indO O (O_reflector O A) (fun _ => A) idmap) _ _). - - intros x; pattern x; apply O_ind_internal. - + intros oa; apply minO_pathsO. + refine (isequiv_adjointify (to (O_reflector A)) + (O_indO (O_reflector A) (fun _ => A) idmap) _ _). + - intros x; pattern x; apply O_ind_easy. + + intros oa; apply inO_pathsO. + intros a; apply ap. - exact (O_indO_beta O (O_reflector O A) (fun _ => A) idmap a). + exact (O_indO_beta (O_reflector A) (fun _ => A) idmap a). - intros a. - exact (O_indO_beta O (O_reflector O A) (fun _ => A) idmap a). + exact (O_indO_beta (O_reflector A) (fun _ => A) idmap a). Defined. (** It seems to be surprisingly hard to show repleteness (without univalence). We basically have to manually develop enough functoriality of [O] and naturality of [to O]. *) - Definition inO_equiv_inO@{u a i j k} (O : Modality@{u a}) (A : Type@{i}) (B : Type@{j}) - (A_inO : In@{u a i} O A) (f : A -> B) (feq : IsEquiv f) - : In@{u a j} O B. + Local Definition inO_equiv_inO_easy (A B : Type) + (A_inO : In_easy A) (f : A -> B) (feq : IsEquiv f) + : In_easy B. Proof. - simple refine (isequiv_commsq (to O A) (to O B) f - (O_ind_internal O A (fun _ => O_reflector O B) _ (fun a => to O B (f a))) _). - - intros; apply O_inO. - - intros a; refine (O_ind_beta_internal@{u a i j k} O A (fun _ => O_reflector O B) _ _ a). + simple refine (isequiv_commsq (to A) (to B) f + (O_ind_easy A (fun _ => O_reflector B) _ (fun a => to B (f a))) _). + - intros; apply O_inO_easy. + - intros a; refine (O_ind_easy_beta A (fun _ => O_reflector B) _ _ a). - apply A_inO. - simple refine (isequiv_adjointify _ - (O_ind_internal O B (fun _ => O_reflector O A) _ (fun b => to O A (f^-1 b))) _ _); + (O_ind_easy B (fun _ => O_reflector A) _ (fun b => to A (f^-1 b))) _ _); intros x. - + apply O_inO. - + pattern x; refine (O_ind_internal O B _ _ _ x); intros. - * apply minO_pathsO. - * simpl; abstract (repeat rewrite O_ind_beta_internal; apply ap, eisretr). - + pattern x; refine (O_ind_internal O A _ _ _ x); intros. - * apply minO_pathsO. - * simpl; abstract (repeat rewrite O_ind_beta_internal; apply ap, eissect). + + apply O_inO_easy. + + pattern x; refine (O_ind_easy B _ _ _ x); intros. + * apply inO_pathsO. + * simpl; abstract (repeat rewrite O_ind_easy_beta; apply ap, eisretr). + + pattern x; refine (O_ind_easy A _ _ _ x); intros. + * apply inO_pathsO. + * simpl; abstract (repeat rewrite O_ind_easy_beta; apply ap, eissect). Defined. - Definition minO_paths@{u a i} (O : Modality@{u a}) - (A : Type@{i}) (A_inO : In@{u a i} O A) (a a' : A) - : In@{u a i} O (a = a'). + Local Definition inO_paths_easy (A : Type) (A_inO : In_easy A) (a a' : A) + : In_easy (a = a'). Proof. - simple refine (inO_equiv_inO O (to O A a = to O A a') _ _ - (@ap _ _ (to O A) a a')^-1 _). - - apply minO_pathsO. + simple refine (inO_equiv_inO_easy (to A a = to A a') _ _ + (@ap _ _ (to A) a a')^-1 _). + - apply inO_pathsO. - refine (@isequiv_ap _ _ _ A_inO _ _). - apply isequiv_inverse. Defined. - (** We don't bother including separatedness for modalities constructed in this way. (We could.) *) - Definition IsSepFor@{u a} (O' O : Modality@{u a}) : Type@{u} - := Empty. - - Definition inO_paths_from_inSepO@{u a i iplus} - (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A) - : In@{u a i} O (x = y) - := Empty_rec sep. - - Definition inSepO_from_inO_paths@{u a i iplus} - (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (op : forall (x y : A), In@{u a i} O (x = y)) - : In@{u a i} O' A - := Empty_rec sep. - -End EasyModalities_to_Modalities. - -(** We now move on to the general theory of modalities. *) - -Module Modalities_Theory (Os : Modalities). - -Export Os. -Module Export Os_ReflectiveSubuniverses - := Modalities_to_ReflectiveSubuniverses Os. -Module Export RSU - := ReflectiveSubuniverses_Theory Os_ReflectiveSubuniverses. - -(** As with reflective subuniverses, we put this in a module so it can be exported separately (and it should be). *) -Module Export Coercions. - Coercion modality_to_reflective_subuniverse - := idmap : Modality -> ReflectiveSubuniverse. -End Coercions. - -(** As promised, we redefine [O_ind] and [O_ind_beta] so that their [In] hypotheses refer to the typeclass [RSU.In]. *) -Definition O_ind@{u a i j e0 e1 e2} {O : Modality@{u a}} - {A : Type@{i}} (B : O A -> Type@{j}) - {B_inO : forall oa, In O (B oa)} - (f : forall a, B (to O A a)) (oa : O A) -: B oa -:= O_ind_internal@{u a e0 e1 e2} O A B B_inO f oa. - -Definition O_ind_beta {O : Modality@{u a}} - {A : Type@{i}} (B : O A -> Type@{j}) - {B_inO : forall oa, In O (B oa)} - (f : forall a : A, B (to O A a)) (a : A) -: @O_ind O A B B_inO f (to O A a) = f a -:= O_ind_beta_internal O A B B_inO f a. - -(** ** The induction principle [O_ind], like most induction principles, is an equivalence. *) -Section OIndEquiv. - Context {fs : Funext} (O : Modality). - - Section OIndEquivData. - - Context {A : Type} (B : O A -> Type) `{forall a, In O (B a)}. - - Global Instance isequiv_O_ind : IsEquiv (O_ind B). - Proof. - apply isequiv_adjointify with (g := fun h => h oD to O A). - - intros h. apply path_forall. - refine (O_ind (fun oa => O_ind B (h oD to O A) oa = h oa) _). - exact (O_ind_beta B (h oD to O A)). - - intros f. apply path_forall. - exact (O_ind_beta B f). - Defined. - - Definition equiv_O_ind - : (forall a, B (to O A a)) <~> (forall oa, B oa) - := Build_Equiv _ _ (O_ind B) _. - - (** Theorem 7.7.7 *) - Definition isequiv_oD_to_O - : IsEquiv (fun (h : forall oa, B oa) => h oD to O A) - := equiv_isequiv equiv_O_ind^-1. - - End OIndEquivData. - - Local Definition isequiv_o_to_O (A B : Type) (B_inO : In O B) - : IsEquiv (fun (h : O A -> B) => h o to O A) - := isequiv_oD_to_O (fun _ => B). - -End OIndEquiv. - -(** Two modalities are the same if they have the same modal types. *) -Class OeqO (O1 O2 : Modality) - := inO_OeqO : forall A, In O1 A <-> In O2 A. - -Global Instance reflexive_OeqO : Reflexive OeqO | 10. -Proof. - intros O A; reflexivity. -Defined. - -Global Instance symmetric_OeqO : Symmetric OeqO | 10. -Proof. - intros O1 O2 O12 A. - specialize (O12 A). - symmetry; assumption. -Defined. - -Global Instance transitive_OeqO : Transitive OeqO | 10. -Proof. - intros O1 O2 O3 O12 O23 A; split. - - intros A1. - apply (@inO_OeqO O2 O3 O23). - apply (@inO_OeqO O1 O2 O12). - exact A1. - - intros A3. - apply (@inO_OeqO O1 O2 O12). - apply (@inO_OeqO O2 O3 O23). - exact A3. -Defined. - -(** Two equivalent modalities have the same connected types. *) -Global Instance isconnected_OeqO {O1 O2 : Modality} `{OeqO O1 O2} - (A : Type) `{IsConnected O1 A} - : IsConnected O2 A. -Proof. - apply isconnected_from_elim. - intros C C2 f. - apply (isconnected_elim O1); apply inO_OeqO; exact _. -Defined. - -(** We prove some useful consequences of [O_ind] that enhance the behavior of reflective subuniverses. *) -Section Enhancements. - Context {O : Modality}. - - (** Corollary 7.5.8: The unit maps [to O A] are connected. *) - Global Instance conn_map_to_O (A : Type) : IsConnMap O (to O A). - Proof. - apply conn_map_from_extension_elim; intros P ? d. - exists (O_ind P d); intros a. - apply O_ind_beta. - Defined. - - (** Corollary 7.7.8, part 2 *) - Global Instance inO_sigma (A:Type) (B:A -> Type) - `{In O A} `{forall a, In O (B a)} - : In O {x:A & B x}. - Proof. - generalize dependent A. - refine (inO_sigma_from_O_ind _ _). - intros A B ? g. - exists (O_ind B g). - apply O_ind_beta. - Defined. - - (** This implies that the composite of modal maps is modal. *) - Global Instance mapinO_compose {A B C : Type} (f : A -> B) (g : B -> C) - `{MapIn O _ _ f} `{MapIn O _ _ g} - : MapIn O (g o f). - Proof. - intros c. - refine (inO_equiv_inO _ (hfiber_compose f g c)^-1). - Defined. - - (** Corollary 7.3.10. (Theorem 7.3.9 is true for any reflective subuniverse.) *) - Corollary equiv_sigma_inO_O {A} `{In O A} (P : A -> Type) - : {x:A & O (P x)} <~> O {x:A & P x}. - Proof. - transitivity (O {x:A & O (P x)}). - - apply equiv_to_O; exact _. - - apply equiv_O_sigma_O. - Defined. + Definition easy_modality : Modality + := Build_Modality In_easy _ inO_equiv_inO_easy O_reflector O_inO_easy to O_ind_easy O_ind_easy_beta inO_paths_easy. -End Enhancements. +End EasyModalities. -(** An enhancement of Corollary 2.29 of CORS: when O (hence also O') is a modality, the map between fibers is not just an O-equivalence but is O-connected. *) -Global Instance conn_map_functor_hfiber `{Univalence} {O O' : Modality} `{IsSepFor O' O} - {Y X : Type} (f : Y -> X) (x : X) - : IsConnMap O (functor_hfiber (fun y => (to_O_natural O' f y)^) x). -Proof. - intros [oy p]. - rewrite <- (inv_V p). - nrefine (isconnected_equiv' O _ - (hfiber_functor_hfiber (to_O_natural O' f) oy x p^) _). - nrefine (isconnected_hfiber_conn_map - (f := (functor_hfiber (to_O_natural O' f) oy)) (x;p^)). - apply conn_map_SepO_inverts. - (* Actually typeclasss search can do the rest by itself, but to help the human reader we show what's going on. (As a byproduct this makes it marginally faster too). *) - apply O_inverts_isconnected. - - apply conn_map_to_O. - - apply conn_map_to_O. -Defined. (** ** The modal factorization system *) @@ -646,175 +607,3 @@ Section ModalFact. Defined. End ModalFact. - -End Modalities_Theory. - -(** ** Restriction of a family of modalities *) - -(** This is just like restriction of reflective subuniverses. *) -Module Type Modalities_Restriction_Data (Os : Modalities). - - Parameter New_Modality : Type2@{u a}. - - Parameter Modalities_restriction - : New_Modality -> Os.Modality. - -End Modalities_Restriction_Data. - -Module Modalities_Restriction - (Os : Modalities) - (Res : Modalities_Restriction_Data Os) -<: Modalities. - - Definition Modality := Res.New_Modality. - - Definition O_reflector (O : Modality@{u a}) - := Os.O_reflector@{u a i} (Res.Modalities_restriction O). - Definition In (O : Modality@{u a}) - := Os.In@{u a i} (Res.Modalities_restriction O). - Definition O_inO (O : Modality@{u a}) - := Os.O_inO@{u a i} (Res.Modalities_restriction O). - Definition to (O : Modality@{u a}) - := Os.to@{u a i} (Res.Modalities_restriction O). - Definition inO_equiv_inO (O : Modality@{u a}) - := Os.inO_equiv_inO@{u a i j k} (Res.Modalities_restriction O). - Definition hprop_inO (H : Funext) (O : Modality@{u a}) - := Os.hprop_inO@{u a i} H (Res.Modalities_restriction O). - Definition O_ind_internal (O : Modality@{u a}) - := Os.O_ind_internal@{u a i j k} (Res.Modalities_restriction O). - Definition O_ind_beta_internal (O : Modality@{u a}) - := Os.O_ind_beta_internal@{u a i j k} (Res.Modalities_restriction O). - Definition minO_paths (O : Modality@{u a}) - := Os.minO_paths@{u a i} (Res.Modalities_restriction O). - Definition IsSepFor@{u a} (O' O : Modality@{u a}) - := @Os.IsSepFor@{u a} (Res.Modalities_restriction@{u a} O') (Res.Modalities_restriction@{u a} O). - Definition inO_paths_from_inSepO@{u a i iplus} (O' O : Modality@{u a}) - := @Os.inO_paths_from_inSepO@{u a i iplus} (Res.Modalities_restriction@{u a} O') (Res.Modalities_restriction@{u a} O). - Definition inSepO_from_inO_paths@{u a i iplus} (O' O : Modality@{u a}) - := @Os.inSepO_from_inO_paths@{u a i iplus} (Res.Modalities_restriction@{u a} O') (Res.Modalities_restriction@{u a} O). - -End Modalities_Restriction. - -(** ** Union of families of modalities *) - -Module Modalities_FamUnion (Os1 Os2 : Modalities) - <: Modalities. - - Definition Modality : Type2@{u a} - := Os1.Modality@{u a} + Os2.Modality@{u a}. - - Coercion Mod_inl := inl : Os1.Modality -> Modality. - Coercion Mod_inr := inr : Os2.Modality -> Modality. - - Definition O_reflector : forall (O : Modality@{u a}), - Type2le@{i a} -> Type2le@{i a}. - Proof. - intros [O|O]; [ exact (Os1.O_reflector@{u a i} O) - | exact (Os2.O_reflector@{u a i} O) ]. - Defined. - - Definition In : forall (O : Modality@{u a}), - Type2le@{i a} -> Type2le@{i a}. - Proof. - intros [O|O]; [ exact (Os1.In@{u a i} O) - | exact (Os2.In@{u a i} O) ]. - Defined. - - Definition O_inO : forall (O : Modality@{u a}) (T : Type@{i}), - In@{u a i} O (O_reflector@{u a i} O T). - Proof. - intros [O|O]; [ exact (Os1.O_inO@{u a i} O) - | exact (Os2.O_inO@{u a i} O) ]. - Defined. - - Definition to : forall (O : Modality@{u a}) (T : Type@{i}), - T -> O_reflector@{u a i} O T. - Proof. - intros [O|O]; [ exact (Os1.to@{u a i} O) - | exact (Os2.to@{u a i} O) ]. - Defined. - - Definition inO_equiv_inO : - forall (O : Modality@{u a}) (T : Type@{i}) (U : Type@{j}) - (T_inO : In@{u a i} O T) (f : T -> U) (feq : IsEquiv f), - In@{u a j} O U. - Proof. - intros [O|O]; [ exact (Os1.inO_equiv_inO@{u a i j k} O) - | exact (Os2.inO_equiv_inO@{u a i j k} O) ]. - Defined. - - Definition hprop_inO - : Funext -> forall (O : Modality@{u a}) (T : Type@{i}), - IsHProp (In@{u a i} O T). - Proof. - intros ? [O|O]; [ exact (Os1.hprop_inO@{u a i} _ O) - | exact (Os2.hprop_inO@{u a i} _ O) ]. - Defined. - - Definition O_ind_internal - : forall (O : Modality@{u a}) - (A : Type2le@{i a}) (B : O_reflector O A -> Type2le@{j a}) - (B_inO : forall oa, In@{u a j} O (B oa)), - (forall a, B (to O A a)) -> forall a, B a. - Proof. - intros [O|O]; [ exact (Os1.O_ind_internal@{u a i j k} O) - | exact (Os2.O_ind_internal@{u a i j k} O) ]. - Defined. - - Definition O_ind_beta_internal - : forall (O : Modality@{u a}) - (A : Type@{i}) (B : O_reflector O A -> Type@{j}) - (B_inO : forall oa, In@{u a j} O (B oa)) - (f : forall a : A, B (to O A a)) (a:A), - O_ind_internal@{u a i j k} O A B B_inO f (to O A a) = f a. - Proof. - intros [O|O]; [ exact (Os1.O_ind_beta_internal@{u a i j k} O) - | exact (Os2.O_ind_beta_internal@{u a i j k} O) ]. - Defined. - - Definition minO_paths - : forall (O : Modality@{u a}) - (A : Type2le@{i a}) (A_inO : In@{u a i} O A) (z z' : A), - In@{u a i} O (z = z'). - Proof. - intros [O|O]; [ exact (Os1.minO_paths@{u a i} O) - | exact (Os2.minO_paths@{u a i} O) ]. - Defined. - - Definition IsSepFor@{u a} - : forall (O' O : Modality@{u a}), Type@{u}. - Proof. - intros [O'|O'] [O|O]; - [ exact (@Os1.IsSepFor@{u a} O' O) - | exact Empty - | exact Empty - | exact (@Os2.IsSepFor@{u a} O' O) ]. - Defined. - - Definition inO_paths_from_inSepO@{u a i iplus} - : forall (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A), - In@{u a i} O (x = y). - Proof. - intros [O'|O'] [O|O]; - [ exact (@Os1.inO_paths_from_inSepO@{u a i iplus} O' O) - | contradiction - | contradiction - | exact (@Os2.inO_paths_from_inSepO@{u a i iplus} O' O) ]. - Defined. - - Definition inSepO_from_inO_paths@{u a i iplus} - : forall (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}), - (forall (x y : A), In@{u a i} O (x = y)) -> In@{u a i} O' A. - Proof. - intros [O'|O'] [O|O]; - [ exact (@Os1.inSepO_from_inO_paths@{u a i iplus} O' O) - | contradiction - | contradiction - | exact (@Os2.inSepO_from_inO_paths@{u a i iplus} O' O) ]. - Defined. - -End Modalities_FamUnion. - -(** For examples of modalities, see the files Notnot, Identity, Nullification, PropositionalFracture, and Localization. *) diff --git a/theories/Modalities/Notnot.v b/theories/Modalities/Notnot.v index 4525ec4d698..d2f2d93c43a 100644 --- a/theories/Modalities/Notnot.v +++ b/theories/Modalities/Notnot.v @@ -9,68 +9,19 @@ Local Open Scope path_scope. (** This is Exercise 7.12 in the book. Note that it is (apparently) *not* accessible unless we assume propositional resizing. *) -(** We are defining only one modality, but it depends on funext. Therefore, we let the family of modalities be the type [Funext]. However, since there is a coercion [O_reflector] from [Modality] to [Funclass], we don't want to simply *define* [Modality] to be [Funext], or else we could start applying [Funext] hypotheses to types and having it act as double negation. - -Instead, we define a [Record] wrapper around it. This is the recommended best practice for all modules (with one exception, see Truncations.v). The constructor of the record should generally be a short name (here [Notnot]) that makes sense as the reflector. *) - -Record Notnot_Modality := Notnot { unNotnot : Funext }. - -Module Notnot_Easy_Modalities <: EasyModalities. - - Definition Modality : Type2@{u a} - := Notnot_Modality. - - Definition O_reflector : Modality@{u a} -> Type@{i} -> Type@{i} - (** We call [not] explicitly with universe annotations so that [O_reflector] has the right number of universe parameters to satisfy the module type. *) - := fun O X => not@{i i} (not@{i i} X). - - Definition to (O : Modality@{u a}) (T : Type@{i}) - : T -> O_reflector@{u a i} O T - := fun x nx => nx x. - - Definition O_indO@{u a i j} (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}) - : (forall a : A, O_reflector@{u a j} O (B (to O A a))) -> - forall z : O_reflector@{u a i} O A, O_reflector@{u a j} O (B z). - Proof. - intros f z nBz. - pose (unNotnot O). (** Access the [Funext] hypothesis *) +Definition NotNot `{Funext} : Modality. +Proof. + snrapply easy_modality. + - intros X; exact (~ (~ X)). + - intros T x nx; exact (nx x). + - intros A B f z nBz. apply z; intros a. - pose proof (hprop_Empty@{i}). - exact (f a (transport (fun x => not@{j j} (B x)) + exact (f a (transport (fun x => ~ (B x)) (path_ishprop _ _) nBz)). - Defined. - - Definition O_indO_beta@{u a i j} (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}) - (f : forall a, O_reflector@{u a j} O (B (to O A a))) (a:A) - : O_indO O A B f (to O A a) = f a. - Proof. - pose (unNotnot O). - pose proof (hprop_Empty@{j}). + - intros A B f a. apply path_ishprop. - Defined. - - Definition minO_pathsO@{u a i} (O : Modality@{u a}) (A : Type@{i}) - (z z' : O_reflector@{u a i} O A) - : IsEquiv@{i i} (to@{u a i} O (z = z')). - Proof. - pose (unNotnot O). - pose proof (hprop_Empty@{i});pose proof (@trunc_hprop@{i}). + - intros A z z'. refine (isequiv_iff_hprop _ _). intros; apply path_ishprop. - Defined. - -End Notnot_Easy_Modalities. - -Module Notnot_Modalities <: Modalities := EasyModalities_to_Modalities Notnot_Easy_Modalities. - -(** After we define any family of modalities or reflective subuniverses, we give a corresponding name to the theory module, generally by postfixing the above-mentioned record constructor with [M] (for "module"). This way, one can either [Import] the theory module (here [NotnotM]) and get access to all the modality functions for the modalities in question, or not import it and access them qualified as [NotnotM.function_name]. *) -Module NotNotM := Modalities_Theory Notnot_Modalities. -Export NotNotM.Coercions. -Export NotNotM.RSU.Coercions. - -(** Finally, we declare a coercion allowing us to use elements of the record wrapper as modalities. *) -Coercion Notnot_Modalities_to_Modalities := idmap - : Notnot_Modality -> Notnot_Modalities.Modality. +Defined. diff --git a/theories/Modalities/Nullification.v b/theories/Modalities/Nullification.v index e6376be6013..c3963ddcba1 100644 --- a/theories/Modalities/Nullification.v +++ b/theories/Modalities/Nullification.v @@ -42,157 +42,46 @@ Definition ooextendable_over_unit@{i j k l m} : ooExtendableAlong_Over (@const A Unit tt) C D ext := fun n => extendable_over_unit n A C D (ext n) (fun c => ext' c n). -(** A basic operation on local generators is the pointwise suspension. *) -Definition susp_nullgen@{a} (S : NullGenerators@{a}) : NullGenerators@{a}. +Definition Nul@{a i} (S : NullGenerators@{a}) : Modality@{i}. Proof. - econstructor; intros i. - exact (Susp@{a a a a a a} (S i)). -Defined. - -(** We define a wrapper, as before. *) -Record Nullification_Modality := Nul { unNul : NullGenerators }. - -Module Nullification_Modalities <: Modalities. - - Definition Modality : Type@{u} := Nullification_Modality@{a}. - (** We use the localization reflective subuniverses for most of the necessary data. *) - Module LocRSU_Data <: ReflectiveSubuniverses_Restriction_Data Localization_ReflectiveSubuniverses. - Definition New_ReflectiveSubuniverse : let enforce := Type@{u'} : Type@{u} in Type@{u} - := Nullification_Modality@{u'}. - Definition ReflectiveSubuniverses_restriction - : New_ReflectiveSubuniverse@{u a} - -> Localization_ReflectiveSubuniverses.ReflectiveSubuniverse@{u a} - := fun O => Loc (null_to_local_generators (unNul O)). - End LocRSU_Data. - - Module LocRSU := ReflectiveSubuniverses_Restriction Localization_ReflectiveSubuniverses LocRSU_Data. - - Module LocRSUTh := ReflectiveSubuniverses_Theory LocRSU. - - Definition O_reflector@{u a i} := LocRSU.O_reflector@{u a i}. - Definition In@{u a i} := LocRSU.In@{u a i}. - Definition O_inO@{u a i} := @LocRSU.O_inO@{u a i}. - Definition to@{u a i} := LocRSU.to@{u a i}. - Definition inO_equiv_inO := @LocRSU.inO_equiv_inO@{u a i j k}. - Definition hprop_inO@{u a i} := LocRSU.hprop_inO@{u a i}. - - Definition O_ind_internal@{u a i j k} (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}) - (B_inO : forall oa : O_reflector@{u a i} O A, In@{u a j} O (B oa)) - (g : forall a : A, B (to@{u a i} O A a)) - : forall x, B x. - Proof. - refine (Localize_ind@{a i j k} - (null_to_local_generators@{a a} (unNul O)) A B g _); intros i. - apply (ooextendable_over_unit@{a i j a k}); intros c. - refine (ooextendable_postcompose - (fun (_:Unit) => B (c tt)) _ _ - (fun u => transport@{i j} B (ap c (path_unit@{a} tt u))) _). - refine (ooextendable_islocal _ i). - apply B_inO. - Defined. - - Definition O_ind_beta_internal (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}) - (B_inO : forall oa : O_reflector O A, In@{u a j} O (B oa)) - (f : forall a : A, B (to O A a)) (a : A) - : O_ind_internal@{u a i j k} O A B B_inO f (to O A a) = f a - := 1. - - Definition minO_paths (O : Modality@{u a}) (A : Type@{i}) - (A_inO : In@{u a i} O A) (z z' : A) - : In@{u a i} O (z = z'). - Proof. - apply (LocRSUTh.inO_paths@{u a i i}); assumption. + simple refine (Build_Modality' (Loc (null_to_local_generators S)) _ _). + - exact _. + - intros A. + (** We take care with universes. *) + simple notypeclasses refine (reflectsD_from_OO_ind@{i} _ _ _). + + intros B B_inO g. + refine (Localize_ind@{a i i i} (null_to_local_generators S) A B g _); intros i. + apply ooextendable_over_unit; intros c. + refine (ooextendable_postcompose@{a a i i i i i i i i} + (fun (_:Unit) => B (c tt)) _ _ + (fun u => transport B (ap c (path_unit@{a} tt u))) _). + refine (ooextendable_islocal _ i). + + reflexivity. + + apply inO_paths@{i i i}. Defined. - (** Just as for general localizations, the separated modality corresponding to a nullification is nullification at the suspension of the generators. The proofs are a little more involved because we have to transfer across the equivalence [Susp Unit <~> Unit]. *) - Definition IsSepFor@{u a} (O' O : Modality@{u a}) : Type@{u} - := paths@{u} (Nul (susp_nullgen (unNul O))) O'. - - Definition inO_paths_from_inSepO@{u a i iplus} - (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A) - : In@{u a i} O (x = y). - Proof. - destruct O as [S]; destruct sep; unfold In, IsLocal in *; intros i; cbn in *. - specialize (A_inO i). - cbn in A_inO. - assert (ee : ooExtendableAlong (functor_susp (fun _:S i => tt)) (fun _ => A)). - { refine (cancelL_ooextendable _ _ (fun _ => tt) _ A_inO). - apply ooextendable_equiv. - apply isequiv_contr_contr. } - assert (e := fst (ooextendable_iff_functor_susp@{a a i i a a a i i iplus i a a a i i iplus i i i i i i i i i i i i iplus i i i iplus i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i} (fun _:S i => tt) _) ee (x,y)). - cbn in e. - refine (ooextendable_postcompose' _ _ _ _ e). - intros b. - symmetry; apply equiv_dp_const. - Defined. - - Definition inSepO_from_inO_paths@{u a i iplus} - (O' O : Modality@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (e : forall (x y : A), In@{u a i} O (x = y)) - : In@{u a i} O' A. - Proof. - destruct O as [S]; destruct sep; unfold In, IsLocal in *; intros i; cbn in *. - apply (ooextendable_compose _ (functor_susp (fun _:S i => tt)) (fun _:Susp Unit => tt)). - 1:apply ooextendable_equiv, isequiv_contr_contr. - apply (ooextendable_iff_functor_susp@{a a i i a a a i i iplus i a a a i i iplus i i i i i i i i i i i i iplus i i i iplus i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i i} (fun _:S i => tt)). - intros [x y]. - refine (ooextendable_postcompose' _ _ _ _ (e x y i)). - intros b. - apply equiv_dp_const. - Defined. - -End Nullification_Modalities. - -(** If you import the following module [NulM], then you can call all the reflective subuniverse functions with a [NullGenerators] as the parameter. *) -Module Import NulM := Modalities_Theory Nullification_Modalities. -(** If you don't import it, then you'll need to write [NulM.function_name]. *) -Export NulM.Coercions. -Export NulM.RSU.Coercions. - -Coercion Nullification_Modality_to_Modality := idmap - : Nullification_Modality -> Modality. - (** And here is the "real" definition of the notation [IsNull]. *) Notation IsNull f := (In (Nul f)). (** ** Nullification and Accessibility *) (** Nullification modalities are accessible, essentially by definition. *) -Module Accessible_Nullification - <: Accessible_Modalities Nullification_Modalities. - - Module Import Os_Theory := Modalities_Theory Nullification_Modalities. - - Definition acc_gen : Modality -> NullGenerators - := unNul. - - Definition inO_iff_isnull (O : Modality@{u a}) (X : Type@{i}) - : iff@{i i i} (In@{u a i} O X) (IsNull (acc_gen O) X) - := (idmap , idmap). - -End Accessible_Nullification. - -Module Import AccNulM := Accessible_Modalities_Theory Nullification_Modalities Accessible_Nullification. - -(** And accessible modalities can be nudged into nullifications. *) - -Module Nudge_Modalities - (Os : Modalities) - (Acc : Accessible_Modalities Os). +Global Instance accmodality_nul (S : NullGenerators) : IsAccModality (Nul S). +Proof. + unshelve econstructor. + - exact S. + - intros; reflexivity. +Defined. - (** Application of modules is still "restricted to paths". *) - Module Data <: Modalities_Restriction_Data Nullification_Modalities. - Definition New_Modality := Os.Modality. - Definition Modalities_restriction - : New_Modality -> Nullification_Modalities.Modality - := Nul o Acc.acc_gen. - End Data. +(** And accessible modalities can be lifted to other universes. *) - Module Nudged <: Modalities - := Modalities_Restriction Nullification_Modalities Data. +Definition lift_accmodality@{a i j} (O : Subuniverse@{i}) `{IsAccModality@{a i} O} + : Modality@{j} + := Nul@{a j} (acc_ngen O). -End Nudge_Modalities. +Global Instance O_eq_lift_accmodality (O : Subuniverse@{i}) `{IsAccModality@{a i} O} + : O <=> lift_accmodality O. +Proof. + split; intros A; apply inO_iff_isnull. +Defined. diff --git a/theories/Modalities/Open.v b/theories/Modalities/Open.v index 3ffb20f1011..247dfc1e483 100644 --- a/theories/Modalities/Open.v +++ b/theories/Modalities/Open.v @@ -9,99 +9,57 @@ Local Open Scope path_scope. (** * Open modalities *) -Record Open_Modality := - Op { funext_Op : Funext ; - unOp : hProp - }. +(** ** Definition *) -Arguments Op {_} _. - -(** ** Open modalities *) - -(** Exercise 7.13(i): Open modalities *) -Module OpenModalities_easy <: EasyModalities. - - Definition Modality : Type@{u} := Open_Modality@{a}. - - Definition O_reflector : Modality@{u a} -> Type@{i} -> Type@{i} - := fun fU X => unOp fU -> X. - - Definition to (O : Modality@{u a}) (T : Type@{i}) : T -> O_reflector@{u a i} O T - := fun x u => x. - - Definition O_indO (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector O A -> Type@{j}) - (f : forall a, O_reflector@{u a j} O (B (to@{u a i} O A a))) - : forall z, O_reflector@{u a j} O (B z). - Proof. - intros z u; pose (funext_Op O). - refine (transport@{i j} B _ (f (z u) u)). +Definition Op `{Funext} (U : hProp) : Modality. +Proof. + snrapply easy_modality. + - intros X; exact (U -> X). + - intros T x; cbn. + exact (fun _ => x). + - cbn; intros A B f z u. + refine (transport B _ (f (z u) u)). apply path_arrow; intros u'. - unfold to; apply ap; apply path_ishprop. - Defined. - - Definition O_indO_beta (O : Modality@{u a}) (A : Type@{i}) - (B : O_reflector@{u a i} O A -> Type@{j}) - (f : forall a, O_reflector@{u a j} O (B (to O A a))) (a : A) - : O_indO O A B f (to O A a) = f a. - Proof. - pose (funext_Op O); apply path_arrow; intros u. + apply ap; apply path_ishprop. + - cbn; intros A B f a. + apply path_arrow; intros u. transitivity (transport B 1 (f a u)); auto with path_hints. apply (ap (fun p => transport B p (f a u))). - transitivity (path_arrow (fun _ => a) (fun _ => a) (@ap10 (unOp O) _ _ _ 1)); + transitivity (path_arrow (fun _ => a) (fun _ => a) (@ap10 U _ _ _ 1)); auto with path_hints. - * apply ap@{i i}. + * apply ap. apply path_forall; intros u'. apply ap_const. * apply eta_path_arrow. - Defined. - - Definition minO_pathsO (O : Modality@{u a}) (A : Type@{i}) - (z z' : O_reflector@{u a i} O A) - : IsEquiv@{i i} (to O (z = z')). - Proof. - pose (fs := funext_Op O); simple refine (isequiv_adjointify _ _ _ _). + - intros A z z'. + srefine (isequiv_adjointify _ _ _ _). * intros f; apply path_arrow; intros u. exact (ap10 (f u) u). * intros f; apply path_arrow; intros u. transitivity (path_arrow z z' (ap10 (f u))). - + unfold to; apply ap@{i i}. + + unfold to; apply ap. apply path_forall; intros u'. apply (ap (fun u0 => ap10 (f u0) u')). apply path_ishprop. + apply eta_path_arrow. * intros p. refine (eta_path_arrow _ _ _). - Defined. - -End OpenModalities_easy. - -Module OpenModalities <: Modalities - := EasyModalities_to_Modalities OpenModalities_easy. - -Module Import OpM := Modalities_Theory OpenModalities. -Export OpM.Coercions. -Export OpM.RSU.Coercions. - -Coercion Open_Modality_to_Modality := - idmap : Open_Modality -> OpenModalities.Modality. +Defined. (** ** The open modality is lex *) -(** Note that unlike the case for closed and topological modalities, we can prove this without univalence (though we do of course need funext). *) -Module Import Lex_OpenModalities := Lex_Modalities_Theory OpenModalities. - -Global Instance lex_open (O : Modality) -: Lex O. +(** Note that unlike most other cases, we can prove this without univalence (though we do of course need funext). *) +Global Instance lex_open `{Funext} (U : hProp) + : Lex (Op U). Proof. - intros A x y. - pose (U := unOp O); pose proof (funext_Op O). - change (Contr (U -> A) -> Contr (U -> (x = y))); intros ?. + apply lex_from_isconnected_paths. + intros A ? x y. assert (uc : U -> Contr A). { intros u. pose (contr_inhabited_hprop U u). - refine (contr_equiv (U -> A) (equiv_contr_forall _)). } + refine (contr_equiv (U -> A) (equiv_contr_forall _)). + apply Ac. } simple refine (Build_Contr _ _ _). - intros u; pose (uc u); exact (center (x=y)). - intros f; apply path_arrow; intros u. @@ -110,42 +68,23 @@ Defined. (** ** The open modality is accessible. *) -Module Accessible_OpenModalities <: Accessible_Modalities OpenModalities. - - Module Import Os_Theory := Modalities_Theory OpenModalities. - - Definition acc_gen - := fun (O : OpenModalities.Modality@{u a}) => - Build_NullGenerators@{a} Unit (fun _ => unOp O). - - Definition inO_iff_isnull' - (O : OpenModalities.Modality@{u a}) (X : Type@{i}) - : iff@{i i i} - (OpenModalities.In@{u a i} O X) - (IsNull (acc_gen O) X). - Proof. - pose (funext_Op O); split. - - intros X_inO u. - apply (equiv_inverse (equiv_ooextendable_isequiv@{a a i i i i i i i i i i i i} _ _)). +Global Instance acc_open `{Funext} (U : hProp) + : IsAccModality (Op U). +Proof. + unshelve econstructor. + - econstructor. + exact (unit_name U). + - intros X; split. + + intros X_inO u. + apply (equiv_inverse (equiv_ooextendable_isequiv _ _)). refine (cancelR_isequiv (fun x (u:Unit) => x)). apply X_inO. - - intros ext; specialize (ext tt). + + intros ext; specialize (ext tt). refine (isequiv_compose (f := (fun x => unit_name x)) - (g := (fun h => h o (@const (unOp O) Unit tt)))). - refine (isequiv_ooextendable@{a a i i i i i i i i i} (fun _ => X) (@const@{i a} (unOp O) Unit tt) ext). - Defined. - - (* Phantom universes (ie not appearing in the term) are produced by some un-annotated functions and foralls. - For instance if we explicitate [(B:=(Unit -> _))] in the [isequiv_compose] use above only one extra is produced. - I can't figure out where the last extra is coming from so we have to manually identify it with [i]. - Since we have to do this manual identification step we might as well use it for the universe from [B] too. - - Phantom universes are pruned by Coq >=8.8 (see Coq/Coq#1033). *) - Definition inO_iff_isnull@{u a i} - := Eval unfold inO_iff_isnull' in ltac:(first [exact inO_iff_isnull'@{u a i i i}|exact inO_iff_isnull'@{u a i}]). - -End Accessible_OpenModalities. + (g := (fun h => h o (@const U Unit tt)))). + refine (isequiv_ooextendable (fun _ => X) (@const U Unit tt) ext). +Defined. (** Thus, arguably a better definition of [Op] would be as a nullification modality, as it would not require [Funext] and would have a judgmental computation rule. However, the above definition is also nice to know, as it doesn't use HITs. We name the other version [Op']. *) -Definition Op' (U : hProp) : Nullification_Modality - := Nul (Build_NullGenerators Unit (fun (_:Unit) => U)). +Definition Op' (U : hProp) : Modality + := Nul (Build_NullGenerators Unit (fun _ => U)). diff --git a/theories/Modalities/ReflectiveSubuniverse.v b/theories/Modalities/ReflectiveSubuniverse.v index 105a68d14c7..501a765e24b 100644 --- a/theories/Modalities/ReflectiveSubuniverse.v +++ b/theories/Modalities/ReflectiveSubuniverse.v @@ -8,7 +8,6 @@ Require Import Tactics.RewriteModuloAssociativity. Local Open Scope nat_scope. Local Open Scope path_scope. - (** * Reflective Subuniverses *) (** ** References *) @@ -20,243 +19,143 @@ Local Open Scope path_scope. - CORS: Christensen, Opie, Rijke, and Scoccola, "Localization in Homotopy Type Theory", https://arxiv.org/abs/1807.04155. *) -(** ** Overview *) - -(** We will define reflective subuniverses using modules. Since modules are one of the more difficult parts of Coq to understand, and the documentation in the reference manual is a bit sparse, we include here a brief introduction to modules. - -For our purposes here, it is appropriate to think of a [Module Type] as analogous to a [Record] type, and a [Module] having that module type (called an "implementation" of it) as analogous to an element of that record type. For instance, instead of - -<< -Record foo := - { bar : Type ; - baz : bar -> Type - }. ->> - -we could write - -<< -Module Type foo. - Parameter bar : Type. - Parameter baz : bar -> Type -End foo. ->> - -and then instead of - -<< -Definition qux : foo - := Build_foo Bool (fun b => if b then Unit else Empty). ->> - -we could write - -<< -Module qux <: foo. - Definition bar : Type - := Bool. - Definition baz : bar -> Type - := fun b => if b then Unit else Empty. -End qux. ->> - -Given these definitions, where we refer to [bar qux] and [baz qux] in the record case, in the module case we would write [qux.bar] and [qux.baz]. However, there are a few essential differences (apart from these syntactic ones). - -Firstly, while elements of records are (like everything else in Coq's type theory) strongly typed, modules are duck-typed. In other words, [qux] is a module of type [foo] simply by virtue of containing fields [bar] and [baz] that have the same types as those declared for the parameters of [foo]; the type declaration [<: foo] only serves to document and enforce this fact. +(** ** Definitions *) -Secondly, modules do not have to be declared to have any type, or they can have more than one type. A module is free to contain as many definitions (and other things such as notations, coercions, instances, etc.) as you like, and to "implement" as many module types as you like. In particular, [qux] could contain additional definitions and it would still be of type [foo]. +(** *** Subuniverses *) -Thirdly, and more importantly, modules are *second-class*: you cannot pass them around as arguments to functions. Nor can you construct them "on the fly"; they can only be defined at top level. However, you can pass a module as an argument to *another module*. For instance, here is a module which takes a module of type [foo] as an argument. +Record Subuniverse@{i} := +{ + In_internal : Type@{i} -> Type@{i} ; + hprop_inO_internal : Funext -> forall (T : Type@{i}), + IsHProp (In_internal T) ; + inO_equiv_inO_internal : forall (T U : Type@{i}) (T_inO : In_internal T) + (f : T -> U) {feq : IsEquiv f}, + In_internal U ; +}. -<< -Module htns (f : foo). - Definition qjkx : Type - := { x : f.bar & f.baz x }. -End htns. ->> +(** Work around Coq bug that fields of records can't be typeclasses. *) +Class In (O : Subuniverse) (T : Type) := in_internal : In_internal O T. -Now if we have a [foo], such as [qux], we can pass it as an argument to [htns] and get a new module (again, only at top level): +(** Being in the subuniverse is a mere predicate (by hypothesis). We include funext in the hypotheses of hprop_inO so that it doesn't have to be assumed in all definitions of (reflective) subuniverses, since in most examples it is required for this and this only. Here we redefine it using the replaced [In]. *) +Global Instance hprop_inO `{Funext} (O : Subuniverse) (T : Type) + : IsHProp (In O T) + := @hprop_inO_internal _ _ T. -<< -Module gcrl := htns qux. ->> - -After this, we can refer to [gcrl.qjkx] and get [{ x : qux.bar & qux.baz x }]. Together with the fact that modules don't need to have a type, this sort of gives us a way to pass a module as an argument to a collection of functions; we can define a module like [htns] which takes a [foo] as an argument and in which we define many functions depending on [foo]; then whenever we want to apply these functions to a [foo] (such as [qux]) we do the application at top-level, as above with [gcrl]. - -Unfortunately, Coq does not allow modules to take elements of ordinary types as arguments either; if you want to pass a [nat], say, as an argument to a module, you have to first wrap the [nat] in another module. You can think of types and module-types as "parallel universes" of types; never the twain shall meet. - -Given these annoying limitations, why would anyone ever use modules instead of records? One reason is that modules are good at (indeed, are more or less designed for) *namespacing*. In particular, it is possible to [Import] a module, so that all of its fields can be accessed without a dot-prefix. In fact, every file in Coq is implicitly its own module, and when you say [Require Import Filename.] you are actually [Import]ing a module. Similarly, modules are used for access control in the private-inductive-types hack that we use to define HITs that compute. - -Another reason to use modules, which is the primary reason we choose to use them here, is that the fields of a module are *individually* universe polymorphic. In other words, in order to define a module of type [foo], as above, you need to give a *polymorphic* definition of [bar] and a *polymorphic* definition of [baz], and the resulting module remembers the polymorphism of each of those fields. By contrast, a definition of an element of a record type may be itself polymorphic, but an individual *instance* of that definition will pertain only to a fixed collection of universes. - -Note that the possibility of individually polymorphic fields practically mandates that modules *must* be second-class. For a polymorphic field involves an implicit quantification over all universes; hence if the record itself were a first-class object, what universe would it live in? A mathematician can think of modules as analogous to the proper classes in NBG set theory: they can be "large" without impacting the consistency strength, *because* we are limited in what we can do with them. - -In the case in point, if a reflective subuniverse were a record, then "a reflective subuniverse" would be a reflective subuniverse of only *one* universe. A polymorphic definition of a particular reflective subuniverse would result in defining related reflective subuniverses of every universe, but the relation *between* these subuniverses would not be specified. In particular, if we have types [X : Type@{i}] and [Y : Type@{j}] in different universes and a map [f : X -> Y], while [Y] is in the subuniverse of [Type@{j}], we could not apply the universal property to extend [f] to a map [O X -> Y], since the universal property asserted for [O@{i} X] would only refer to maps with target also in [Type@{i}]. This is at best annoying; for instance, it means that we couldn't define, say, [Trunc_functor] by using [O_functor] and then proceed to apply it to maps between types in different universes (which turns out to be necessary sometimes). At worst, such as when trying to prove that the universe of modal types for a lex modalities is itself modal, this approach seems more or less unworkable. - -Therefore, we choose to make a reflective subuniverse a module. This means that in order to define "a reflective subuniverse", you have to give a *polymorphic* definition of the reflector, the universal property, etc. In particular, the universal property must be polymorphic enough to allow the situation with [X : Type@{i}] and [Y : Type@{j}] considered above. - -There are some issues involving this choice that must be addressed. One of them is that when implementing a polymorphic module types, Coq is *very* strict about matching up the polymorphism. Specifically, each [Definition] in the implementing module must have *exactly* the same number of universe parameters as the corresponding [Parameter] in the module type, and all the constraints in the former must be implied by those in the latter. This ensures that the implementation is "at least as polymorphic" as the specification. - -Now normally, a polymorphic definition will end up with many more universes than it needs, and we have little control over how many those are. Therefore, in order to have a chance of ensuring that our implementations of module types match up in polymorphism, we almost always need to add explicit universe annotations in order to control how many universe parameters they end up with. This is slightly annoying, but fortunately it only needs to be dealt with when *defining* a particular reflective subuniverse; to users the polymorphism should be invisible and automatic. - -This also means it is important that we know exactly how many universe parameters each field of our module types is *expected* to take. It would be nice if Coq had a feature for declaring (and verifying) the universe parameters of a definition in the same way that we declare the type parameters. In the absence of this (requested at https://coq.inria.fr/bugs/show_bug.cgi?id=3818), we write [Check foo@{a b c}.] after the definition of [foo] to declare that [foo] takes three universe parameters. Note that this will fail with an [Error] unless [foo] does in fact take exactly three universe parameters. - -Another issue that must be dealt with is the fact, mentioned above, that a module cannot be parametrized over an ordinary type. However, it frequently happens that we do want to define a family of reflective subuniverses, e.g. the n-truncation modalities for all [n : trunc_index], or the open and closed modalities for all [U : hProp]. The solution we choose is for our basic [Module Type] to represent not a *single* reflective subuniverse, but an entire *family* of them, parametrized by some type. This can be regarded as analogous to how when doing mathematics relative to a base topos, the correct notion of "large category" is an *indexed category* (a.k.a. fibration), which comes with a basic notion of "[I]-indexed family of objects" for all [I] in the base topos. -*) +(** We assumed repleteness of the subuniverse in the definition. Of course, with univalence this would be automatic, but we include it as a hypothesis since most of the theory of reflective subuniverses and modalities doesn't need univalence, and most or all examples can be shown to be replete without using univalence. Here we redefine it using the replaced [In]. *) +Definition inO_equiv_inO {O : Subuniverse} (T : Type) {U : Type} + `{T_inO : In O T} (f : T -> U) `{IsEquiv T U f} + : In O U + := @inO_equiv_inO_internal O T U T_inO f _. -Module Type ReflectiveSubuniverses. +Definition inO_equiv_inO' {O : Subuniverse} + (T : Type) {U : Type} `{In O T} (f : T <~> U) + : In O U + := inO_equiv_inO T f. - (** As mentioned above, an implementation of this module type is a *family* of reflective subuniverses, indexed by the below type [ReflectiveSubuniverse]. If we just wrote [ReflectiveSubuniverse : Type], then it would end up parametrized by one universe, but in many examples the natural definition of the parametrizing type involves also a smaller universe, which would cause problems with Coq's strict polymorphism enforcement for module type implementations. Thus, we use [Type2] instead, which takes two universe parameters. *) - Parameter ReflectiveSubuniverse@{u a} : Type2@{u a}. +(** The universe of types in the subuniverse *) +Definition Type_@{i j} (O : Subuniverse@{i}) : Type@{j} + := @sig@{j i} Type@{i} (fun (T : Type@{i}) => In O T). - (** The universe parameters occurring in the definitions here play one of the following roles, which we indicate consistently by the same letters. +Coercion TypeO_pr1 O (T : Type_ O) := @pr1 Type (In O) T. - - [u] is the size of the parametrizing type [ReflectiveSubuniverse] (and, later, also [Modality]). - - [a] is the size of smaller type-data occurring in that type, such as the family of generators for a localization. This generally must be strictly smaller than [u]. - - [i] is the size of a type that we are reflecting or testing to be in the subuniverse. This is generally at least as big as [a]. - - [j] is the size of a type that we are eliminating into (out of a type in [i]) with a universal property. Also generally at least as big as [a]. - - [k] is a universe at least as large as both [i] and [j], in which statements about types in both of them can live. - - [iplus] is a universe that is strictly larger than [i]. *) - - Parameter O_reflector@{u a i} : forall (O : ReflectiveSubuniverse@{u a}), - Type2le@{i a} -> Type2le@{i a}. - - (** For reflective subuniverses (and hence also modalities), it will turn out that [In O T] is equivalent to [IsEquiv (O_unit T)]. We could define the former as the latter, and it would simplify some of the general theory. However, in many examples there is a "more basic" definition of [In O] which is equivalent, but not definitionally identical, to [IsEquiv (O_unit T)]. Thus, including [In O] as data makes more things turn out to be judgmentally what we would expect. *) - Parameter In@{u a i} : forall (O : ReflectiveSubuniverse@{u a}), - Type2le@{i a} -> Type2le@{i a}. - - Parameter O_inO@{u a i} : forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - In@{u a i} O (O_reflector@{u a i} O T). - - Parameter to@{u a i} : forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - T -> O_reflector@{u a i} O T. - - Parameter inO_equiv_inO@{u a i j k} : - forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}) (U : Type@{j}) - (T_inO : In@{u a i} O T) (f : T -> U) (feq : IsEquiv f), - (** We add an extra universe parameter that's at least as large as both [i] and [j]. This seems to be necessary for the proof of repleteness in some examples, such as easy modalities. *) - let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - In@{u a j} O U. - - (** In most examples, [Funext] is necessary to prove that the predicate of being in the subuniverse is an hprop. To avoid needing to assume [Funext] as a global hypothesis when constructing such examples, and since [Funext] is often not needed for any of the rest of the theory, we add it as a hypothesis to this specific field. *) - Parameter hprop_inO@{u a i} - : Funext -> forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - IsHProp (In@{u a i} O T). - - (** We express the universal property using the representation [ooExtendableAlong] of precomposition equivalences. This has the advantage that it avoids the funext redexes that otherwise infect the theory, thereby simplifying the proofs and proof terms. We never have to worry about whether we have a path between functions or a homotopy; we use only homotopies, with no need for [ap10] or [path_arrow] to mediate. Furthermore, the data in [ooExtendableAlong] are all special cases of the induction principle of a modality. Thus, all the theorems we prove about reflective subuniverses will, when interpreted for a modality (coerced as above to a reflective subuniverse), reduce definitionally to "the way we would have proved them directly for a modality". *) - Parameter extendable_to_O@{u a i j k} - : forall (O : ReflectiveSubuniverse@{u a}) {P : Type2le@{i a}} {Q : Type2le@{j a}} {Q_inO : In@{u a j} O Q}, - ooExtendableAlong@{i i j k} (to O P) (fun _ => Q). - - (** As defined in CORS, for a reflective subuniverse [O], the [O]-separated types are those whose identity types are in [O]. These form another reflective subuniverse [O'], but in general it might not be in the same family. Thus, instead of considering [O] -> [O'] as an operation on reflective subuniverses, we consider it a relation between them; [IsSepFor O' O] says that [O'] is the [O]-separated subuniverse. In addition, we can't just *define* this relation to say that the types in [O'] are the [O]-separated ones, because we want that to be true universe-polymorphically. Thus, we assert the relation [IsSepFor] as part of the data of a family of reflective subuniverses, along with two more parameters saying that the types in [O'] are the [O]-separated ones at all universe levels. *) - Parameter IsSepFor@{u a} : forall (O' O : ReflectiveSubuniverse@{u a}), Type@{u}. - - (** These two have an extra universe parameter [iplus] that's strictly bigger than [i], because our current proofs of them in the case of localization seem to produce such a universe. (It's unclear how necessary that is.) *) - Parameter inO_paths_from_inSepO@{u a i iplus} - : forall (O' O : ReflectiveSubuniverse@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A), - let gt := (Type1@{i} : Type@{iplus}) in - In@{u a i} O (x = y). - - Parameter inSepO_from_inO_paths@{u a i iplus} - : forall (O' O : ReflectiveSubuniverse@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}), - let gt := (Type1@{i} : Type@{iplus}) in - (forall (x y : A), In@{u a i} O (x = y)) -> In@{u a i} O' A. - -End ReflectiveSubuniverses. - - -(** We now begin a parametrized module to incorporate most of the theory of reflective subuniverses. Thus, after defining a particular family of reflective subuniverses, you can apply this module and [Import] it to get all of the theory. (Some suggested naming conventions for these modules can be found in Modality.v.) *) -Module ReflectiveSubuniverses_Theory (Os : ReflectiveSubuniverses). -Export Os. - -(** Membership in the subuniverse is a typeclass. *) -Existing Class In. - -(** The type of types in the subuniverse *) -Definition Type_@{u a j i} (O : ReflectiveSubuniverse@{u a}) : Type@{j} - := @sig@{j i} Type@{i} (fun (T : Type@{i}) => In@{u a i} O T). - -(** Before going on, we declare some coercions in a module, so that they can be imported separately. In fact, this submodule should be exported by any file that defines a reflective subuniverse. *) -Module Export Coercions. - - (** We allow the name of a subuniverse or modality to be used as the name of its reflector. This means that when defining a particular example, you should generally put the parametrizing family in a wrapper, so that you can notate the subuniverse as parametrized by, rather than identical to, its parameter. See Modality.v, Truncations.v, and Localization.v for examples. *) - Coercion O_reflector : ReflectiveSubuniverse >-> Funclass. - - (** Obviously, every element of [Type_ O] is a type. *) - Coercion TypeO_pr1 O (T : Type_ O) := @pr1 Type (In O) T. - -End Coercions. - -(** We assumed repleteness of the subuniverse in the definition. Of course, with univalence this would be automatic, but we include it as a hypothesis since most of the theory of reflective subuniverses and modalities doesn't need univalence, and most or all examples can be shown to be replete without using univalence. *) -Arguments inO_equiv_inO {O} T {U} {_} f {_}. +(** The second component of [TypeO] is unique. *) +Definition path_TypeO@{i j} {fs : Funext} O (T T' : Type_ O) (p : T.1 = T'.1) + : T = T' + := path_sigma_hprop@{j i j} T T' p. -(** Being in the subuniverse is a mere predicate (by hypothesis) *) -Global Existing Instance hprop_inO. +Definition equiv_path_TypeO@{i j} {fs : Funext} O (T T' : Type_ O) + : (paths@{j} T.1 T'.1) <~> (T = T') + := equiv_path_sigma_hprop@{j i j j} T T'. -(** [O T] is always in the subuniverse (by hypothesis). This needs a universe annotation to become sufficiently polymorphic. *) -Arguments O_inO {O} T. +(** Types in [TypeO] are always in [O]. *) +Global Instance inO_TypeO {O : Subuniverse} (A : Type_ O) : In O A + := A.2. + +(** *** Reflections *) + +(** A pre-reflection is a map to a type in the subuniverse. *) +Class PreReflects@{i} (O : Subuniverse@{i}) (T : Type@{i}) := +{ + O_reflector : Type@{i} ; + O_inO : In O O_reflector ; + to : T -> O_reflector ; +}. + +Arguments O_reflector O T {_}. +Arguments to O T {_}. +Arguments O_inO {O} T {_}. Global Existing Instance O_inO. -Existing Class IsSepFor. -Global Existing Instance inO_paths_from_inSepO. +(** It is a reflection if it has the requisite universal property. *) +Class Reflects@{i} (O : Subuniverse@{i}) (T : Type@{i}) + `{PreReflects@{i} O T} := +{ + extendable_to_O : forall {Q : Type@{i}} {Q_inO : In O Q}, + ooExtendableAlong (to O T) (fun _ => Q) +}. + +Arguments extendable_to_O O {T _ _ Q Q_inO}. + +(** Here's a modified version that applies to types in possibly-smaller universes without collapsing those universes to [i]. *) +Definition extendable_to_O'@{i j k} (O : Subuniverse@{i}) (T : Type@{j}) + `{Reflects O T} {Q : Type@{k}} {Q_inO : In O Q} + : ooExtendableAlong@{j i k i} (to O T) (fun _ => Q). +Proof. + assert (e := @extendable_to_O O T _ _ Q Q_inO). + apply (lift_ooextendablealong@{i j j i i i i k k i i}). + exact e. +Defined. + +Record ReflectiveSubuniverse@{i} := +{ + rsu_subuniv : Subuniverse@{i} ; + rsu_prereflects : forall (T : Type@{i}), PreReflects rsu_subuniv T ; + rsu_reflects : forall (T : Type@{i}), Reflects rsu_subuniv T ; +}. + +Coercion rsu_subuniv : ReflectiveSubuniverse >-> Subuniverse. +Global Existing Instance rsu_prereflects. +Global Existing Instance rsu_reflects. + +(** We allow the name of a subuniverse or modality to be used as the name of its reflector. This means that when defining a particular example, you should generally put the parametrizing family in a wrapper, so that you can notate the subuniverse as parametrized by, rather than identical to, its parameter. See Modality.v, Truncations.v, and Localization.v for examples. *) +Definition rsu_reflector (O : ReflectiveSubuniverse) (T : Type) : Type + := O_reflector O T. + +Coercion rsu_reflector : ReflectiveSubuniverse >-> Funclass. + +(** *** Recursion principles *) -(** The second component of [TypeO] is unique. Without a universe annotation, for some reason this collapses [j=u] and [i=a]. *) -Definition path_TypeO@{u a j i} {fs : Funext} O (T T' : Type_@{u a j i} O) (p : T.1 = T'.1) - : T = T' - := path_sigma_hprop T T' p. +(** We now extract the recursion principle and the restricted induction principles for paths. *) +Section ORecursion. + Context {O : Subuniverse} {P Q : Type} {Q_inO : In O Q} `{Reflects O P}. -Definition equiv_path_TypeO@{u a j i} {fs : Funext} O (T T' : Type_@{u a j i} O) -: (paths@{j} T.1 T'.1) <~> (T = T') -:= equiv_path_sigma_hprop@{j i j j} T T'. + Definition O_rec (f : P -> Q) + : O_reflector O P -> Q + := (fst (extendable_to_O O 1%nat) f).1. -(** Types in [TypeO] are always in [O]. *) -Global Instance inO_TypeO {O : ReflectiveSubuniverse} (A : Type_ O) -: In O A -:= A.2. + Definition O_rec_beta (f : P -> Q) (x : P) + : O_rec f (to O P x) = f x + := (fst (extendable_to_O O 1%nat) f).2 x. -Definition extendable_to_O (O : ReflectiveSubuniverse) - {P Q : Type} {Q_inO : In O Q} -: ooExtendableAlong (to O P) (fun _ => Q) - := @extendable_to_O O P Q Q_inO. + Definition O_indpaths (g h : O_reflector O P -> Q) + (p : g o to O P == h o to O P) + : g == h + := (fst (snd (extendable_to_O O 2) g h) p).1. -(** We now extract the recursion principle and the restricted induction principles for paths. *) -Section ORecursion. - Context {O : ReflectiveSubuniverse}. - - Definition O_rec {P Q : Type} {Q_inO : In O Q} - (f : P -> Q) - : O P -> Q - := (fst (extendable_to_O O 1%nat) f).1. - - Definition O_rec_beta {P Q : Type} {Q_inO : In O Q} - (f : P -> Q) (x : P) - : O_rec f (to O P x) = f x - := (fst (extendable_to_O O 1%nat) f).2 x. - - Definition O_indpaths {P Q : Type} {Q_inO : In O Q} - (g h : O P -> Q) (p : g o to O P == h o to O P) - : g == h - := (fst (snd (extendable_to_O O 2) g h) p).1. - - Definition O_indpaths_beta {P Q : Type} {Q_inO : In O Q} - (g h : O P -> Q) (p : g o (to O P) == h o (to O P)) (x : P) - : O_indpaths g h p (to O P x) = p x - := (fst (snd (extendable_to_O O 2) g h) p).2 x. - - Definition O_ind2paths {P Q : Type} {Q_inO : In O Q} - {g h : O P -> Q} (p q : g == h) + Definition O_indpaths_beta (g h : O_reflector O P -> Q) + (p : g o (to O P) == h o (to O P)) (x : P) + : O_indpaths g h p (to O P x) = p x + := (fst (snd (extendable_to_O O 2) g h) p).2 x. + + Definition O_ind2paths {g h : O_reflector O P -> Q} (p q : g == h) (r : p oD (to O P) == q oD (to O P)) - : p == q - := (fst (snd (snd (extendable_to_O O 3) g h) p q) r).1. + : p == q + := (fst (snd (snd (extendable_to_O O 3) g h) p q) r).1. - Definition O_ind2paths_beta {P Q : Type} {Q_inO : In O Q} - {g h : O P -> Q} (p q : g == h) + Definition O_ind2paths_beta {g h : O_reflector O P -> Q} (p q : g == h) (r : p oD (to O P) == q oD (to O P)) (x : P) - : O_ind2paths p q r (to O P x) = r x - := (fst (snd (snd (extendable_to_O O 3) g h) p q) r).2 x. + : O_ind2paths p q r (to O P x) = r x + := (fst (snd (snd (extendable_to_O O 3) g h) p q) r).2 x. (** Clearly we can continue indefinitely as needed. *) @@ -291,12 +190,11 @@ Global Instance isequiv_O_rec_to_O `{Funext} (** We now prove a bunch of things about an arbitrary reflective subuniverse. *) Section Reflective_Subuniverse. - Universes Ou Oa. - Context (O : ReflectiveSubuniverse@{Ou Oa}). + Context (O : ReflectiveSubuniverse). (** Functoriality of [O_rec] homotopies *) Definition O_rec_homotopy {P Q : Type} `{In O Q} (f g : P -> Q) (pi : f == g) - : O_rec f == O_rec g. + : O_rec (O := O) f == O_rec g. Proof. apply O_indpaths; intro x. etransitivity. @@ -307,12 +205,11 @@ Section Reflective_Subuniverse. Defined. (** If [T] is in the subuniverse, then [to O T] is an equivalence. *) - Definition isequiv_to_O_inO@{u a i} (T : Type@{i}) `{In@{u a i} O T} : IsEquiv@{i i} (to O T). + Definition isequiv_to_O_inO (T : Type) `{In O T} : IsEquiv (to O T). Proof. - (** Using universe annotations to reduce superfluous universes *) - pose (g := O_rec@{u a i i i i i} idmap). + pose (g := O_rec idmap : O T -> T). refine (isequiv_adjointify (to O T) g _ _). - - refine (O_indpaths@{u a i i i i i} (to O T o g) idmap _). + - refine (O_indpaths (to O T o g) idmap _). intros x. apply ap. apply O_rec_beta. @@ -340,7 +237,7 @@ Section Reflective_Subuniverse. Definition O_functor_compose {A B C : Type} (f : A -> B) (g : B -> C) : (O_functor (g o f)) == (O_functor g) o (O_functor f). Proof. - apply O_indpaths; intros x. + srapply O_indpaths; intros x. refine (to_O_natural (g o f) x @ _). transitivity (O_functor g (to O B (f x))). - symmetry. exact (to_O_natural g (f x)). @@ -509,9 +406,9 @@ Section Reflective_Subuniverse. := ap_O_path_universe' (Build_Equiv _ _ f _). (** Postcomposition respects [O_rec] *) - Definition O_rec_postcompose {A B C : Type} `{In O B} {C_inO : In O C} + Definition O_rec_postcompose {A B C : Type@{i}} `{In O B} {C_inO : In O C} (f : A -> B) (g : B -> C) - : g o O_rec f == O_rec (g o f). + : g o O_rec (O := O) f == O_rec (O := O) (g o f). Proof. refine (O_indpaths _ _ _); intros x. transitivity (g (f x)). @@ -523,10 +420,6 @@ Section Reflective_Subuniverse. Section Replete. - Definition inO_equiv_inO'@{i j k} (T : Type@{i}) {U : Type@{j}} `{In@{Ou Oa i} O T} (f : T <~> U) - : In@{Ou Oa j} O U - := inO_equiv_inO@{Ou Oa i j k} T f. - (** An equivalent formulation of repleteness is that a type lies in the subuniverse as soon as its unit map is an equivalence. *) Definition inO_isequiv_to_O (T:Type) : IsEquiv (to O T) -> In O T @@ -542,7 +435,7 @@ Section Reflective_Subuniverse. Defined. (** Thus, [T] is in a subuniverse as soon as [to O T] admits a retraction. *) - Definition inO_to_O_retract@{i | Oa <= i} (T:Type@{i}) (mu : O T -> T) + Definition inO_to_O_retract (T:Type) (mu : O T -> T) : Sect (to O T) mu -> In O T. Proof. unfold Sect; intros H. @@ -557,7 +450,7 @@ Section Reflective_Subuniverse. Section OInverts. - (** The maps that are inverted by the reflector. Note that this notation is NOT GLOBAL, it only exists in this section. *) + (** The maps that are inverted by the reflector. Note that this notation is NOT (yet) GLOBAL; it only exists in this section. *) Local Notation O_inverts f := (IsEquiv (O_functor f)). Global Instance O_inverts_O_unit (A : Type) @@ -586,18 +479,44 @@ Section Reflective_Subuniverse. Definition isequiv_O_rec_O_inverts {A B : Type} `{In O B} (f : A -> B) `{O_inverts f} - : IsEquiv (O_rec f). + : IsEquiv (O_rec (O := O) f). Proof. apply isequiv_O_inverts. - refine (cancelR_isequiv (O_functor (to O A))). - refine (isequiv_homotopic (O_functor (O_rec f o to O A)) + nrefine (cancelR_isequiv (O_functor (to O A))); [ exact _ | ]. + nrefine (isequiv_homotopic (O_functor (O_rec f o to O A)) (O_functor_compose _ _)). refine (isequiv_homotopic (O_functor f) (O_functor_homotopy _ _ (fun x => (O_rec_beta f x)^))). Defined. - (** If a map is inverted by the representable functor [? -> Z] for all [O]-modal types [Z], then it is inverted by [O]. First we prove a version that doesn't require funext. *) - Definition O_inverts_by_yoneda' {A : Type@{i}} {B : Type@{j}} (f : A -> B) + (** If [f] is inverted by [O], then mapping out of it into any modal type is an equivalence. First we prove a version not requiring funext. For use in [O_inverts_O_leq] below, we allow the types [A], [B], and [Z] to perhaps live in smaller universes than the one [i] on which our subuniverse lives. This the first half of Lemma 1.23 of RSS. *) + Definition ooextendable_O_inverts@{a b z i} + {A : Type@{a}} {B : Type@{b}} (f : A -> B) `{O_inverts f} + (Z : Type@{z}) `{In@{i} O Z} + : ooExtendableAlong@{a b z i} f (fun _ => Z). + Proof. + refine (cancelL_ooextendable@{a b i z i i i i i} _ _ (to O B) _ _). + 1:exact (extendable_to_O'@{i b z} O B). + refine (ooextendable_homotopic _ (O_functor f o to O A) _ _). + 1:apply to_O_natural. + refine (ooextendable_compose _ (to O A) (O_functor f) _ _). + - srapply ooextendable_equiv. + - exact (extendable_to_O'@{i a z} O A). + Defined. + + (** And now the funext version *) + Definition isequiv_precompose_O_inverts `{Funext} + {A B : Type} (f : A -> B) `{O_inverts f} + (Z : Type) `{In O Z} + : IsEquiv (fun g:B->Z => g o f). + Proof. + srapply (equiv_extendable_isequiv 0). + exact (ooextendable_O_inverts f Z 2). + Defined. + + (** Conversely, if a map is inverted by the representable functor [? -> Z] for all [O]-modal types [Z], then it is inverted by [O]. As before, first we prove a version that doesn't require funext. *) + Definition O_inverts_from_extendable + {A : Type@{i}} {B : Type@{j}} (f : A -> B) (** Without the universe annotations, the result ends up insufficiently polymorphic. *) (e : forall (Z:Type@{k}), In O Z -> ExtendableAlong@{i j k l} 2 f (fun _ => Z)) : O_inverts f. @@ -617,41 +536,14 @@ Section Reflective_Subuniverse. Defined. (** And the version with funext. Use it with universe parameters [i j k l lplus l l l l]. *) - Definition O_inverts_by_yoneda `{Funext} - {A : Type@{i}} {B : Type@{j}} (f : A -> B) - (e : forall (Z:Type@{k}), In@{Ou Oa k} O Z -> - IsEquiv@{l l} (fun g:B->Z => g o f)) + Definition O_inverts_from_isequiv_precompose `{Funext} + {A B : Type} (f : A -> B) + (e : forall (Z:Type), In O Z -> IsEquiv (fun g:B->Z => g o f)) : O_inverts f. Proof. - apply O_inverts_by_yoneda'. + apply O_inverts_from_extendable. intros Z ?. - apply (equiv_extendable_isequiv 0 _ _)^-1. - rapply e. - Defined. - - (** The converse is also true. This the first half of Lemma 1.23 of RSS. *) - Definition ooextendable_O_inverts - {A B : Type} (f : A -> B) `{O_inverts f} - (Z : Type) `{In O Z} - : ooExtendableAlong f (fun _ => Z). - Proof. - refine (cancelL_ooextendable _ _ (to O B) _ _). - 1:srapply extendable_to_O. - refine (ooextendable_homotopic _ (O_functor f o to O A) _ _). - 1:apply to_O_natural. - refine (ooextendable_compose _ (to O A) (O_functor f) _ _). - 2:srapply extendable_to_O. - srapply ooextendable_equiv. - Defined. - - (** And now the funext version *) - Definition isequiv_precompose_O_inverts `{Funext} - {A B : Type} (f : A -> B) `{O_inverts f} - (Z : Type) `{In O Z} - : IsEquiv (fun g:B->Z => g o f). - Proof. - srapply (equiv_extendable_isequiv 0). - exact (ooextendable_O_inverts f Z 2). + rapply ((equiv_extendable_isequiv 0 _ _)^-1%equiv). Defined. (** This property also characterizes the types in the subuniverse, which is the other half of Lemma 1.23. *) @@ -659,7 +551,7 @@ Section Reflective_Subuniverse. (E : forall (A : Type@{i}) (B : Type@{j}) (f : A -> B) (Oif : O_inverts f), ooExtendableAlong f (fun _ => Z)) - : In@{Ou Oa k} O Z. + : In O Z. Proof. pose (EZ := fst (E Z (O Z) (to O Z) _ 1%nat) idmap). exact (inO_to_O_retract _ EZ.1 EZ.2). @@ -774,7 +666,7 @@ Section Reflective_Subuniverse. Proof. intro H. pose (ev := fun x => (fun (f:(forall x, (B x))) => f x)). - pose (zz := fun x:A => O_rec (ev x)). + pose (zz := fun x:A => O_rec (O := O) (ev x)). apply inO_to_O_retract with (mu := fun z => fun x => zz x z). intro phi. unfold zz, ev; clear zz; clear ev. @@ -794,7 +686,7 @@ Section Reflective_Subuniverse. : In O (A*B). Proof. apply inO_to_O_retract with - (mu := fun X => (@O_rec _ (A * B) A _ fst X , O_rec snd X)). + (mu := fun X => (@O_rec _ (A * B) A _ _ _ fst X , O_rec snd X)). intros [a b]; apply path_prod; simpl. - exact (O_rec_beta fst (a,b)). - exact (O_rec_beta snd (a,b)). @@ -808,7 +700,7 @@ Section Reflective_Subuniverse. Proof. refine (equiv_uncurry A B C oE _). refine (_ oE (equiv_uncurry (O A) (O B) C)^-1). - refine (equiv_o_to_O _ A (B -> C) oE _); simpl. + refine (equiv_o_to_O O A (B -> C) oE _); simpl. apply equiv_postcompose'. exact (equiv_o_to_O _ B C). Defined. @@ -879,97 +771,6 @@ Section Reflective_Subuniverse. apply concat_pV_p ). Defined. - (** ** Dependent sums *) - (** Theorem 7.7.4, (ii) => (i) *) - Definition inO_sigma_from_O_ind@{i j} - : (forall (A:Type@{i}) (B: O_reflector@{Ou Oa i} O A -> Type@{j}) `{forall a, In@{Ou Oa j} O (B a)} - (g : forall (a:A), (B (to@{Ou Oa i} O A a))), - sig@{i j} (fun f : forall (z:O_reflector@{Ou Oa i} O A), (B z) => forall a:A, f (to@{Ou Oa i} O A a) = g a)) - -> - (forall (A:Type@{i}) (B:A -> Type@{j}) `{In@{Ou Oa i} O A} `{forall a, In@{Ou Oa j} O (B a)}, - (In@{Ou Oa j} O (sig@{i j} (fun x:A => B x)))). - Proof. - intros H A B ? ?. - pose (h := fun x => @O_rec@{Ou Oa i i i j i} _ ({x:A & B x}) A _ pr1 x). - pose (p := (fun z => O_rec_beta pr1 z) - : h o (to O _) == pr1). - pose (g := fun z => (transport B ((p z)^) z.2)). - simpl in *. - specialize (H ({x:A & B x}) (fun x => B (h x)) _ g). - destruct H as [f q]. - apply inO_to_O_retract@{i} with (mu := fun w => (h w; f w)). - intros [x1 x2]. - simple refine (path_sigma B _ _ _ _); simpl. - - apply p. - - refine (ap _ (q (x1;x2)) @ _). - unfold g; simpl. exact (transport_pV B _ _). - Defined. - - (** For (i) => (ii) we first prove a "local" version, that if a *particular* sigma-type is in [O] then it admits extensions. *) - Definition extension_from_inO_sigma@{i j k} - {A:Type@{i}} (B: (O A) -> Type@{j}) - {H : In@{Ou Oa k} O (sig@{i j} (fun z:O A => B z))} - (g : forall x, B (to O A x)) - : ExtensionAlong@{i i j k k} (to O A) B g. - Proof. - set (Z := sigT B : Type@{k}) in *. - pose (g' := (fun a:A => (to O A a ; g a)) : A -> Z). - pose (f' := O_rec@{Ou Oa i k i k k} g'). - pose (eqf := (O_rec_beta g') : f' o to O A == g'). - pose (eqid := O_indpaths@{Ou Oa i i i i i} - (pr1 o f') idmap (fun x => ap@{k i} pr1 (eqf x))). - exists (fun z => transport B (eqid z) ((f' z).2)). - intros a. unfold eqid. - refine (_ @ pr2_path (O_rec_beta g' a)). - refine (ap@{i k} (fun p => transport B p (O_rec g' (to O A a)).2) _). - srapply O_indpaths_beta. - Defined. - - (** We then deduce the general version from this. Note that although here we see three universe parameters, after the section closes this definition ends up with five universe parameters [Ou Oa i j k]. *) - Definition O_ind_from_inO_sigma@{i j k} - (H : forall (A:Type@{i}) (B:A -> Type@{j}) - {A_inO : In@{Ou Oa i} O A} - `{forall a, In@{Ou Oa j} O (B a)}, - (In@{Ou Oa k} O (sig@{i j} (fun x:A => B x)))) - (A:Type@{i}) (B: (O A) -> Type@{j}) `{forall a, In@{Ou Oa j} O (B a)} - (g : forall (a:A), (B (to O A a))) - := @extension_from_inO_sigma@{i j k} A B (H (O A) B) g. - - (** In fact, we can enhance [extension_from_inO_sigma] to a local version of [extendable_to_O], as stated in CORS Proposition 2.8 (but our version avoids funext by using [ooExtendableAlong], as usual). *) - Definition extendable_from_inO_sigma - {A:Type@{i}} (B: (O A) -> Type@{j}) - {H : In O (sig (fun z:O A => B z))} - : ooExtendableAlong@{i i j k} (to O A) B. - Proof. - intros n; generalize dependent A. - induction n as [|n IHn]; intros; [ exact tt | cbn ]. - refine (extension_from_inO_sigma B , _). - intros h k; apply IHn. - set (Z := sigT B) in *. - pose (W := sigT (fun a => B a * B a)). - refine (inO_equiv_inO' (Pullback (A := W) (fun a:O A => (a;(h a,k a))) - (fun z:Z => (z.1;(z.2,z.2)))) _). - { srapply inO_pullback. - exact (inO_equiv_inO' _ (equiv_sigprod_pullback B B)^-1). } - unfold Pullback. - (** The rest is just contracting a couple of based path spaces. It seems like it should be less work than this. *) - apply equiv_functor_sigma_id; intros z; cbn. - refine (_ oE equiv_functor_sigma_id _). - 2:intros; symmetry; apply equiv_path_sigma. - refine (_ oE (equiv_sigma_assoc _ _)^-1%equiv); cbn. - refine (_ oE equiv_functor_sigma_id _). - 2:intros; apply equiv_sigma_symm. - refine (_ oE (equiv_sigma_assoc' _ _)). - refine (_ oE equiv_contr_sigma _); cbn. - refine (_ oE equiv_functor_sigma_id _). - 2:{ intros; symmetry; etransitivity; revgoals. - - apply equiv_path_prod. - - apply equiv_sigma_prod0. } - refine (_ oE equiv_sigma_assoc' _ _). - refine (_ oE equiv_contr_sigma _); cbn. - apply equiv_path_inverse. - Defined. - (** ** Fibers *) Global Instance inO_hfiber {A B : Type} `{In O A} `{In O B} @@ -1032,11 +833,11 @@ Section Reflective_Subuniverse. (** Naively it might seem that we need closure under Sigmas (hence a modality) to deduce closure under [Equiv], but in fact the above closure under fibers is sufficient. This appears as part of the proof of Proposition 2.18 of CORS. For later use, we try to reduce the number of universe parameters (but we don't completely control them all). *) Global Instance inO_equiv `{Funext} (A : Type@{i}) (B : Type@{j}) - `{In@{Ou Oa i} O A} `{In@{Ou Oa j} O B} - : In@{Ou Oa k} O (A <~> B). + `{In O A} `{In O B} + : In O (A <~> B). Proof. - refine (inO_equiv_inO@{Ou Oa k k k} _ (issig_equiv@{i j k} A B)). - refine (inO_equiv_inO@{Ou Oa k k k} _ (equiv_functor_sigma equiv_idmap@{k} + refine (inO_equiv_inO _ (issig_equiv@{i j k} A B)). + refine (inO_equiv_inO _ (equiv_functor_sigma equiv_idmap@{k} (fun f => equiv_biinv_isequiv@{i j k l} f))). transparent assert (c : (prod@{k k} (A->B) (prod@{k k} (B->A) (B->A)) -> prod@{k k} (A -> A) (B -> B))). { intros [f [g h]]; exact (h o f, f o g). } @@ -1062,12 +863,11 @@ Section Reflective_Subuniverse. apply (path_sigma' _ 1); cbn. refine (_ @ eta_path_prod p); apply ap011; apply eta_path_arrow. Defined. - Check inO_equiv@{i j k l k k}. (** k is max(i,j), while k < l *) (** ** Paths *) - Definition inO_paths@{i j} (S : Type@{i}) {S_inO : In@{Ou Oa i} O S} (x y : S) - : In@{Ou Oa i} O (x=y). + Definition inO_paths@{i j} (S : Type@{i}) {S_inO : In O S} (x y : S) + : In O (x=y). Proof. simple refine (inO_to_O_retract@{i} _ _ _); intro u. - assert (p : (fun _ : O (x=y) => x) == (fun _=> y)). @@ -1473,7 +1273,7 @@ Notation O_inverts O f := (IsEquiv (O_functor O f)). Question: is there a definition of connectedness (say, for n-types) that neither blows up the universe level, nor requires HIT's? *) (** We give annotations to reduce the number of universe parameters. *) -Class IsConnected (O : ReflectiveSubuniverse@{u a}) (A : Type@{i}) +Class IsConnected (O : ReflectiveSubuniverse@{i}) (A : Type@{i}) := isconnected_contr_O : Contr@{i} (O A). Global Existing Instance isconnected_contr_O. @@ -1503,7 +1303,7 @@ Section ConnectedTypes. Definition isconnected_elim {A : Type} `{IsConnected O A} (C : Type) `{In O C} (f : A -> C) : NullHomotopy f. Proof. - set (ff := @O_rec O _ _ _ f). + set (ff := @O_rec O _ _ _ _ _ f). exists (ff (center _)). intros a. symmetry. refine (ap ff (contr (to O _ a)) @ _). @@ -1564,28 +1364,28 @@ Section ConnectedTypes. Defined. (** Here's another way of stating the universal property for mapping out of connected types into modal ones. *) - Definition extendable_const_isconnected_inO@{Ou Oa i j k l m} (n : nat) + Definition extendable_const_isconnected_inO (n : nat) (** Work around https://coq.inria.fr/bugs/show_bug.cgi?id=3811 *) - (A : Type@{i}) {conn_A : IsConnected@{Ou Oa i} O A} - (C : Type@{j}) `{In@{Ou Oa j} O C} - : ExtendableAlong@{i i j j} n (@const@{i i} A Unit tt) (fun _ => C). + (A : Type@{i}) {conn_A : IsConnected@{i} O A} + (C : Type@{j}) `{In O C} + : ExtendableAlong n (@const@{i i} A Unit tt) (fun _ => C). Proof. generalize dependent C; simple_induction n n IHn; intros C ?; [ exact tt | split ]. - intros f. - exists (fun _ : Unit => (isconnected_elim@{i j k j j i} C f).1); intros a. - symmetry; apply ((isconnected_elim@{i j k j j i} C f).2). + exists (fun _ : Unit => (isconnected_elim C f).1); intros a. + symmetry; apply ((isconnected_elim C f).2). - intros h k. - refine (extendable_postcompose'@{i i j j j j l l l l} n _ _ _ _ (IHn (h tt = k tt) (inO_paths@{Ou Oa j m} _ _ _ _))). + refine (extendable_postcompose' n _ _ _ _ (IHn (h tt = k tt) (inO_paths _ _ _ _))). intros []; apply equiv_idmap. Defined. - Definition ooextendable_const_isconnected_inO@{Ou Oa i j k} - (A : Type@{i}) `{IsConnected@{Ou Oa i} O A} - (C : Type@{j}) `{In@{Ou Oa j} O C} + Definition ooextendable_const_isconnected_inO + (A : Type@{i}) `{IsConnected@{i} O A} + (C : Type@{j}) `{In O C} : ooExtendableAlong (@const A Unit tt) (fun _ => C) - := fun n => extendable_const_isconnected_inO@{Ou Oa i j k k j} n A C. + := fun n => extendable_const_isconnected_inO n A C. Definition isequiv_const_isconnected_inO `{Funext} {A : Type} `{IsConnected O A} (C : Type) `{In O C} @@ -1730,12 +1530,11 @@ End ModalMaps. (** Connectedness of a map can again be defined in two equivalent ways: by connectedness of its fibers (as types), or by the lifting property/elimination principle against truncated types. We use the former; the equivalence with the latter is given below in [conn_map_elim], [conn_map_comp], and [conn_map_from_extension_elim]. *) -Class IsConnMap (O : ReflectiveSubuniverse@{u a}) - {A : Type@{i}} {B : Type@{j}} (f : A -> B) +Class IsConnMap (O : ReflectiveSubuniverse@{i}) + {A : Type@{i}} {B : Type@{i}} (f : A -> B) := isconnected_hfiber_conn_map (** The extra universe [k] is >= max(i,j). *) - : forall b:B, IsConnected@{u a k} O (hfiber@{i j} f b). -Check IsConnMap@{u a i j k}. + : forall b:B, IsConnected@{i} O (hfiber@{i i} f b). Global Existing Instance isconnected_hfiber_conn_map. @@ -1958,7 +1757,7 @@ Section ConnectedMaps. (* Lemma 7.5.10: A map to a type in [O] exhibits its codomain as the [O]-reflection of its domain if it is [O]-connected. (The converse is true if and only if [O] is a modality.) *) Definition isequiv_O_rec_conn_map {A B : Type} `{In O B} (f : A -> B) `{IsConnMap O _ _ f} - : IsEquiv (O_rec f). + : IsEquiv (O_rec (O := O) f). Proof. refine (isequiv_adjointify _ (conn_map_elim f (fun _ => O A) (to O A)) _ _). - intros x. pattern x. @@ -2056,332 +1855,156 @@ Section ConnectedMaps. End ConnectedMaps. +(** ** Containment of (reflective) subuniverses *) + +(** One subuniverse is contained in another if every [O1]-modal type is [O2]-modal. We define this parametrized by three universes: [O1] and [O2] are reflective subuniverses of [Type@{i1}] and [Type@{i2}] respectively, and the relation says that all types in [Type@{j}] that [O1]-modal are also [O2]-modal. This implies [j <= i1] and [j <= i2], of course. The most common application is when [i1 = i2 = j], but it's sometimes useful to talk about a subuniverse of a larger universe agreeing with a subuniverse of a smaller universe on the smaller universe. *) +Class O_leq@{i1 i2 j} (O1 : Subuniverse@{i1}) (O2 : Subuniverse@{i2}) + := inO_leq : forall (A : Type@{j}), In O1 A -> In O2 A. + +Arguments inO_leq O1 O2 {_} A _. + +Declare Scope subuniverse_scope. +Notation "O1 <= O2" := (O_leq O1 O2) : subuniverse_scope. +Open Scope subuniverse_scope. + +Global Instance reflexive_O_leq : Reflexive O_leq | 10. +Proof. + intros O A ?; assumption. +Defined. + +Global Instance transitive_O_leq : Transitive O_leq | 10. +Proof. + intros O1 O2 O3 O12 O23 A ?. + rapply (@inO_leq O2 O3). + rapply (@inO_leq O1 O2). +Defined. + +(** This implies that every [O2]-connected type is [O1]-connected, and similarly for maps and equivalences. We give universe annotations so that [O1] and [O2] don't have to be on the same universe, but we do have to have [i1 <= i2] for this statement. When [i2 <= i1] it seems that the statement might not be true unless the RSU on the larger universe is accessibly extended from the smaller one; see [Localization.v]. *) +Definition isconnected_O_leq@{i1 i2} + (O1 : ReflectiveSubuniverse@{i1}) (O2 : ReflectiveSubuniverse@{i2}) `{O_leq@{i1 i2 i1} O1 O2} + (A : Type@{i1}) `{IsConnected O2 A} + : IsConnected O1 A. +Proof. + apply isconnected_from_elim. + intros C C1 f. + apply (isconnected_elim O2); srapply inO_leq; exact _. +Defined. + +(** This one has the same universe constraint [i1 <= i2]. *) +Definition conn_map_O_leq@{i1 i2} + (O1 : ReflectiveSubuniverse@{i1}) (O2 : ReflectiveSubuniverse@{i2}) `{O_leq@{i1 i2 i1} O1 O2} + {A B : Type@{i1}} (f : A -> B) `{IsConnMap O2 A B f} + : IsConnMap O1 f. +Proof. + (** We could prove this by applying [isconnected_O_leq] fiberwise, but unless we were very careful that would collapse the two universes [i1] and [i2]. So instead we just give an analogous direct proof. *) + apply conn_map_from_extension_elim. + intros P P_inO g. + rapply (extension_conn_map_elim O2). + intros b; rapply (inO_leq O1). +Defined. + +(** This is Lemma 2.12(i) in CORS, again with the same universe constraint [i1 <= i2]. *) +Definition O_inverts_O_leq@{i1 i2} + (O1 : ReflectiveSubuniverse@{i1}) (O2 : ReflectiveSubuniverse@{i2}) `{O_leq@{i1 i2 i1} O1 O2} + {A B : Type@{i1}} (f : A -> B) `{O_inverts O2 f} + : O_inverts O1 f. +Proof. + apply O_inverts_from_extendable@{i1 i1 i1 i1 i1}; intros Z Z_inO. + pose (inO_leq O1 O2 Z _). + apply (lift_extendablealong@{i1 i1 i1 i1 i1 i1 i2 i1 i1 i2 i1}). + apply (ooextendable_O_inverts O2); exact _. +Defined. + + +(** ** Equality of (reflective) subuniverses *) + +(** Two subuniverses are the same if they have the same modal types. The universe parameters are the same as for [O_leq]: [O1] and [O2] are reflective subuniverses of [Type@{i1}] and [Type@{i2}], and the relation says that they agree when restricted to [Type@{j}], where [j <= i1] and [j <= i2]. *) +Class O_eq@{i1 i2 j} (O1 : Subuniverse@{i1}) (O2 : Subuniverse@{i2}) := +{ + O_eq_l : O_leq@{i1 i2 j} O1 O2 ; + O_eq_r : O_leq@{i2 i1 j} O2 O1 ; +}. + +Global Existing Instances O_eq_l O_eq_r. + +Notation "O1 <=> O2" := (O_eq O1 O2) (at level 70) : subuniverse_scope. + +Definition issig_O_eq O1 O2 : _ <~> O_eq O1 O2 := ltac:(issig). + +Global Instance reflexive_O_eq : Reflexive O_eq | 10. +Proof. + intros; split; reflexivity. +Defined. + +Global Instance transitive_O_eq : Transitive O_eq | 10. +Proof. + intros O1 O2 O3; split; refine (transitivity (y := O2) _ _). +Defined. + +Global Instance symmetric_O_eq : Symmetric O_eq | 10. +Proof. + intros O1 O2 [? ?]; split; assumption. +Defined. + +Definition issig_subuniverse : _ <~> Subuniverse := ltac:(issig). + +Definition equiv_path_subuniverse `{Univalence} (O1 O2 : Subuniverse) + : (O1 <=> O2) <~> (O1 = O2). +Proof. + refine (_ oE (issig_O_eq O1 O2)^-1). + revert O1 O2; refine (equiv_path_along_equiv issig_subuniverse _). + cbn; intros O1 O2. + refine (equiv_path_sigma_hprop O1 O2 oE _). + destruct O1 as [O1 [O1h ?]]; destruct O2 as [O2 [O2h ?]]; cbn. + refine (equiv_path_arrow _ _ oE _). + srapply (equiv_iff_hprop). + - srapply trunc_sigma; unfold O_leq; exact _. + - intros [h k] A; specialize (h A); specialize (k A); cbn in *. + apply path_universe_uncurried, equiv_iff_hprop; assumption. + - intros h; split; intros A e; specialize (h A); cbn in *. + 1:rewrite <- h. + 2:rewrite h. + all:exact e. +Defined. + +(** It should also be true that if [O1] and [O2] are reflective subuniverses, then [O1 <=> O2] is equivalent to [O1 = O2 :> ReflectiveSubuniverse]. Probably [contr_typeof_O_unit] should be useful in proving that. *) + +(** Reflections into one subuniverse are also reflections into an equal one. Unfortunately these almost certainly can't be [Instance]s for fear of infinite loops, since [<=>] is reflexive. *) +Definition prereflects_O_leq + (O1 O2 : Subuniverse) `{O1 <= O2} + (A : Type) `{PreReflects O1 A} + : PreReflects O2 A. +Proof. + unshelve econstructor. + - exact (O_reflector O1 A). + - rapply (inO_leq O1 O2). + - exact (to O1 A). +Defined. + +Definition reflects_O_eq + (O1 O2 : Subuniverse) `{O1 <=> O2} + (A : Type) `{Reflects O1 A} + : @Reflects O2 A (prereflects_O_leq O1 O2 A). +Proof. + constructor; intros B B_inO2. + pose proof (inO_leq O2 O1 _ B_inO2). + apply (extendable_to_O O1). +Defined. -(** ** Separated subuniverses *) - -Section Separated. - Universes Ou Oa. - Context {O' O : ReflectiveSubuniverse@{Ou Oa}} {sep : IsSepFor@{Ou Oa} O' O}. - - Global Instance inSepO_inO {A : Type} `{In O A} : In O' A. - Proof. - rapply inSepO_from_inO_paths. - Defined. - - Global Instance inSepO_hprop {A : Type} {hp : IsHProp A} : In O' A. - Proof. - rapply inSepO_from_inO_paths. - Defined. - - (** Proposition 2.18 of CORS. Since we have universe-polymorphic RSUs, we can state it in the simpler form "[Type_ O] is in [O']". Note the universe parameters: [Type_@{Ou Oa j i}] is a subtype of [Type@{i}], which is an element of [Type@{j}] (so that [i B) (equiv_compose'@{i j j} _ _)). - - refine (equiv_path_TypeO@{Ou Oa j i} O (A;_) (B;_)). - - apply equiv_path_universe. - Defined. - - (** Lemma 2.21 of CORS *) - Global Instance inSepO_sigma {X : Type} {P : X -> Type} `{In O' X} `{forall x, In O (P x)} - : In O' (sig P). - Proof. - rapply inSepO_from_inO_paths. - 1:exact _. - intros [x p] [y q]. - refine (inO_equiv_inO' O _ (equiv_path_sigma P (x;p) (y;q))). - Defined. - - (** Proposition 2.22 of CORS (in funext-free form). Universe parameters are [i j k kplus kplus kplus]. *) - Definition extendable_toSepO_inO - {X : Type@{i}} (P : O' X -> Type@{j}) `{forall x, In@{Ou Oa j} O (P x)} - : ooExtendableAlong@{i i j k} (to O' X) P. - Proof. - rapply extendable_from_inO_sigma@{Ou Oa i j k kplus k k k kplus k k kplus}. - Defined. - - (** And now the version with funext. Universe parameters are [i j k kplus kplus kplus]. *) - Definition isequiv_toSepO_inO `{Funext} - {X : Type@{i}} (P : O' X -> Type@{j}) `{forall x, In@{Ou Oa j} O (P x)} - : IsEquiv@{k k} (fun g : (forall y, P y) => g o to O' X) - := isequiv_ooextendable@{i i j k k k k k k k k} _ _ (extendable_toSepO_inO P). - - Definition equiv_toSepO_inO@{i j k kplus} `{Funext} - {X : Type@{i}} (P : O' X -> Type@{j}) `{forall x, In O (P x)} - : (forall y, P y) <~> (forall x, P (to O' X x)) - := Build_Equiv _ _ _ (isequiv_toSepO_inO@{i j k kplus kplus kplus} P). - - (** Proposition 2.26 of CORS *) - Definition path_SepO {X : Type@{i}} (x y : X) - : O (x = y) -> (to O' X x = to O' X y) - := O_rec (O := O) (@ap X (O' X) (to O' X) x y). - - Global Instance isequiv_path_SepO `{Univalence} {X : Type@{i}} (x y : X) - : IsEquiv (path_SepO x y). - Proof. - (* First we define a family [P : O' X -> Type] of [O]-modal types that extends [O (x = ?) : X -> Type], using the fact that the universe of [O]-modal types is [O']-modal. [Pb] is its non-definitional computation rule that we will have to transport across. *) - pose (Oxeq := fun y : X => (O (x = y); O_inO (x = y)) : Type_ O). - assert (Pb := fun y => (O_rec_beta@{Ou Oa i j j i j} Oxeq y)..1). - set (P := O_rec Oxeq) in *; subst Oxeq; cbn in *. - (* We will first construct an equivalence [P y' <~> (to O' X x = y')] by a version of encode-decode, then show that when [y' = to O' X y] this equivalence is homotopic to [path_SepO]. *) - srapply isequiv_homotopic'. - { (* First we generalize from [y:X] to [y':O' X]. *) - refine (_ oE equiv_path _ _ (Pb y)^). - generalize (to O' X y) as y'; clear y. - (* Now by a standard result, it suffices to show that [P] is a pointed predicate whose total space is contractible. *) - apply equiv_path_from_contr. - (* The point is just the image of [idpath x], but we have to transport it across [Pb] to put it in the right type. *) - 1:exact (transport idmap (Pb x)^ (to O _ idpath)). - (* The same point supplies the center of contraction. *) - exists (to O' X x; transport idmap (Pb x)^ (to O _ idpath)). - intros [y' p]; revert y' p. - (* To construct a contraction, we first observe that it suffices to return to the case [y:X]. *) - pose (H1 := @inSepO_sigma); pose (H2 := @inO_paths_from_inSepO); pose (H3 := @inO_forall); - refine ((extension_from_inO_sigma O' _ _).1); clear H1 H2 H3. - intros y. - (* Now we have to transport [p] again across the computation rule [Pb]. *) - equiv_intro (equiv_path _ _ (Pb y)^) p; revert p; cbn. - (* We claim that once again, it suffices to assume [p:x=y]. *) - refine ((extension_from_inO_sigma O _ _).1). - { refine (inO_equiv_inO' _ (to O' X x = to O' X y) _). - (* First we change the equality in a sigma-type to a pair of equalities. *) - refine (equiv_functor_sigma_id _ oE _); cbn. - 1:intros p; apply equiv_path_sigma. - (* Now we rearrange things to see that the two extra components are equivalent to a contractible based path space. *) - cbn. refine (equiv_sigma_symm _ oE _). - refine ((equiv_sigma_contr _)^-1%equiv). intros p. - refine (@contr_equiv' _ _ _ (@contr_basedpaths (O (x = y)) _)). - apply equiv_functor_sigma_id; intros q; cbn. - exact (equiv_moveL_transport_V idmap (Pb y) _ q). } - (* Now it's easy by path induction. *) - intros p; destruct p. - reflexivity. } - (* Finally, since [path_SepO] is the unique map making a commutative triangle with [to O' X], to show that the above equivalence is homotopic to [path_SepO] it suffices to show that it also makes that triangle commute. *) - { unfold equiv_path_from_contr. - apply O_indpaths; intros p. - refine (_ @ (O_rec_beta _ _)^). - (* In fact it's easier to show that the *inverse* of this equivalence makes the triangle commute in the other direction. *) - apply moveR_equiv_V. - (* Which is true by a simple path induction, since that inverse was in turn defined by a path induction. *) - destruct p; cbn. - reflexivity. } - Defined. - - Definition equiv_path_SepO `{Univalence} {A : Type} (x y : A) - : O (x = y) <~> (to O' A x = to O' A y) - := Build_Equiv _ _ _ (isequiv_path_SepO x y). - - Global Instance conn_map_toSepO@{i} {X : Type@{i} } - : IsConnMap O (to O' X). - Proof. - apply conn_map_from_extension_elim@{Ou Oa i i i i}. - intros Q Q_inO h. - exact (fst (extendable_toSepO_inO@{i i i Ou Ou Ou} _ 1%nat) h). - Defined. - - (** A strengthening of Lemma 2.27 of CORS. We provide a different proof that doesn't require [Funext]. *) - Global Instance conn_map_functor_sigma_toSepO@{i j} - {X : Type@{i} } (P : O' X -> Type@{j}) - : IsConnMap O (functor_sigma@{i j i j} (Q := P) (to O' X) (fun x => idmap)). - Proof. - srapply conn_map_functor_sigma@{Ou Oa i i i i i i i i i i i i i i Ou i i i i i i i Ou i i}. - Defined. - - Definition equiv_functor_sigma_to_SepO@{i j} - {X : Type@{i} } (P : O' X -> Type@{j}) - : (O {x : X & P (to O' X x)}) <~> (O {ox : O' X & P ox}) - := Build_Equiv _ _ (O_functor O (functor_sigma@{i j i j} _ _)) _. - - (** Corollary 2.29 of CORS *) - Global Instance SepO_inverts_functor_hfiber `{Univalence} - {Y X : Type@{i} } (f : Y -> X) (x : X) - : O_inverts O (functor_hfiber@{i i i i} (fun y => (to_O_natural O' f y)^) x). - Proof. - srapply isequiv_homotopic'. - - unfold hfiber. - refine (_ oE (equiv_inverse (equiv_O_sigma_O O _))). - refine (equiv_functor_sigma_to_SepO@{i i} _ oE _). - apply equiv_O_functor. - apply equiv_functor_sigma_id; intros y; cbn. - refine (_ oE equiv_path_SepO (f y) x). - apply equiv_concat_l, to_O_natural. - - apply O_indpaths. - intros [y p]; cbn. - abstract ( - rewrite O_rec_beta; - rewrite !(to_O_natural O _ _); - apply ap; refine (path_sigma' _ 1 _); cbn; - unfold path_SepO; - rewrite O_rec_beta, inv_V; reflexivity - ). - Defined. - - Definition equiv_SepO_functor_hfiber `{Univalence} - {Y X : Type@{i} } (f : Y -> X) (x : X) - := Build_Equiv _ _ _ (SepO_inverts_functor_hfiber f x). - - (** Proposition 2.30 of CORS *) - (* Making this an [Instance] breaks typeclass inference in various places. *) - Definition conn_map_SepO_inverts `{Univalence} {Y X : Type@{i}} (f : Y -> X) - `{O_inverts O' f} - : IsConnMap@{Ou Oa i i i} O f. - Proof. - intros x. - refine (contr_equiv' (O (hfiber (O_functor O' f) (to O' X x))) _). - exact ((equiv_SepO_functor_hfiber f x)^-1%equiv). - Defined. - -End Separated. - - -End ReflectiveSubuniverses_Theory. - -(** ** Restriction of a family of reflective subuniverses *) - -(** Recall that an implementation of [ReflectiveSubuniverses] is a family of reflective subuniverses indexed by the type [ReflectiveSubuniverse]. Sometimes we want to consider only a subfamily of a known one, or more generally a restriction of such a family along a function. The second-class nature of modules makes this a bit of a pain to construct, but we can do it. *) -Module Type ReflectiveSubuniverses_Restriction_Data (Os : ReflectiveSubuniverses). - - Parameter New_ReflectiveSubuniverse@{u a} : Type2@{u a}. - - Parameter ReflectiveSubuniverses_restriction@{u a} - : New_ReflectiveSubuniverse@{u a} -> Os.ReflectiveSubuniverse@{u a}. - -End ReflectiveSubuniverses_Restriction_Data. - -Module ReflectiveSubuniverses_Restriction - (Os : ReflectiveSubuniverses) - (Res : ReflectiveSubuniverses_Restriction_Data Os) -<: ReflectiveSubuniverses. - - Definition ReflectiveSubuniverse := Res.New_ReflectiveSubuniverse. - - Definition O_reflector@{u a i} (O : ReflectiveSubuniverse@{u a}) - := Os.O_reflector@{u a i} (Res.ReflectiveSubuniverses_restriction O). - Definition In@{u a i} (O : ReflectiveSubuniverse@{u a}) - := Os.In@{u a i} (Res.ReflectiveSubuniverses_restriction O). - Definition O_inO@{u a i} (O : ReflectiveSubuniverse@{u a}) - := Os.O_inO@{u a i} (Res.ReflectiveSubuniverses_restriction O). - Definition to@{u a i} (O : ReflectiveSubuniverse@{u a}) - := Os.to@{u a i} (Res.ReflectiveSubuniverses_restriction O). - Definition inO_equiv_inO@{u a i j k} (O : ReflectiveSubuniverse@{u a}) - := Os.inO_equiv_inO@{u a i j k} (Res.ReflectiveSubuniverses_restriction O). - Definition hprop_inO@{u a i} (H : Funext) (O : ReflectiveSubuniverse@{u a}) - := Os.hprop_inO@{u a i} H (Res.ReflectiveSubuniverses_restriction O). - Definition extendable_to_O@{u a i j k} (O : ReflectiveSubuniverse@{u a}) - := @Os.extendable_to_O@{u a i j k} (Res.ReflectiveSubuniverses_restriction@{u a} O). - Definition IsSepFor@{u a} (O' O : ReflectiveSubuniverse@{u a}) - := @Os.IsSepFor@{u a} (Res.ReflectiveSubuniverses_restriction@{u a} O') (Res.ReflectiveSubuniverses_restriction@{u a} O). - Definition inO_paths_from_inSepO@{u a i iplus} (O' O : ReflectiveSubuniverse@{u a}) - := @Os.inO_paths_from_inSepO@{u a i iplus} (Res.ReflectiveSubuniverses_restriction@{u a} O') (Res.ReflectiveSubuniverses_restriction@{u a} O). - Definition inSepO_from_inO_paths@{u a i iplus} (O' O : ReflectiveSubuniverse@{u a}) - := @Os.inSepO_from_inO_paths@{u a i iplus} (Res.ReflectiveSubuniverses_restriction@{u a} O') (Res.ReflectiveSubuniverses_restriction@{u a} O). - -End ReflectiveSubuniverses_Restriction. -(** ** Union of families of reflective subuniverses *) - -Module ReflectiveSubuniverses_FamUnion - (Os1 Os2 : ReflectiveSubuniverses) -<: ReflectiveSubuniverses. - - Definition ReflectiveSubuniverse@{u a} : Type2@{u a} - := Os1.ReflectiveSubuniverse@{u a} + Os2.ReflectiveSubuniverse@{u a}. - - Coercion RSU_inl := inl : Os1.ReflectiveSubuniverse -> ReflectiveSubuniverse. - Coercion RSU_inr := inr : Os2.ReflectiveSubuniverse -> ReflectiveSubuniverse. - - Definition O_reflector@{u a i} : forall (O : ReflectiveSubuniverse@{u a}), - Type2le@{i a} -> Type2le@{i a}. - Proof. - intros [O|O]; [ exact (Os1.O_reflector@{u a i} O) - | exact (Os2.O_reflector@{u a i} O) ]. - Defined. - - Definition In@{u a i} : forall (O : ReflectiveSubuniverse@{u a}), - Type2le@{i a} -> Type2le@{i a}. - Proof. - intros [O|O]; [ exact (Os1.In@{u a i} O) - | exact (Os2.In@{u a i} O) ]. - Defined. - - Definition O_inO@{u a i} - : forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - In@{u a i} O (O_reflector@{u a i} O T). - Proof. - intros [O|O]; [ exact (Os1.O_inO@{u a i} O) - | exact (Os2.O_inO@{u a i} O) ]. - Defined. - - Definition to@{u a i} : forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - T -> O_reflector@{u a i} O T. - Proof. - intros [O|O]; [ exact (Os1.to@{u a i} O) - | exact (Os2.to@{u a i} O) ]. - Defined. - - Definition inO_equiv_inO@{u a i j k} : - forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}) (U : Type@{j}) - (T_inO : In@{u a i} O T) (f : T -> U) (feq : IsEquiv f), - In@{u a j} O U. - Proof. - intros [O|O]; [ exact (Os1.inO_equiv_inO@{u a i j k} O) - | exact (Os2.inO_equiv_inO@{u a i j k} O) ]. - Defined. - - Definition hprop_inO@{u a i} - : Funext -> forall (O : ReflectiveSubuniverse@{u a}) (T : Type@{i}), - IsHProp (In@{u a i} O T). - Proof. - intros ? [O|O]; [ exact (Os1.hprop_inO@{u a i} _ O) - | exact (Os2.hprop_inO@{u a i} _ O) ]. - Defined. - - Definition extendable_to_O@{u a i j k} - : forall (O : ReflectiveSubuniverse@{u a}) {P : Type2le@{i a}} {Q : Type2le@{j a}} {Q_inO : In@{u a j} O Q}, - ooExtendableAlong@{i i j k} (to O P) (fun _ => Q). - Proof. - intros [O|O]; [ exact (@Os1.extendable_to_O@{u a i j k} O) - | exact (@Os2.extendable_to_O@{u a i j k} O) ]. - Defined. - - Definition IsSepFor@{u a} - : forall (O' O : ReflectiveSubuniverse@{u a}), Type@{u}. - Proof. - intros [O'|O'] [O|O]; - [ exact (@Os1.IsSepFor@{u a} O' O) - | exact Empty - | exact Empty - | exact (@Os2.IsSepFor@{u a} O' O) ]. - Defined. - - Definition inO_paths_from_inSepO@{u a i iplus} - : forall (O' O : ReflectiveSubuniverse@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}) (A_inO : In@{u a i} O' A) (x y : A), - In@{u a i} O (x = y). - Proof. - intros [O'|O'] [O|O]; - [ exact (@Os1.inO_paths_from_inSepO@{u a i iplus} O' O) - | contradiction - | contradiction - | exact (@Os2.inO_paths_from_inSepO@{u a i iplus} O' O) ]. - Defined. - - Definition inSepO_from_inO_paths@{u a i iplus} - : forall (O' O : ReflectiveSubuniverse@{u a}) (sep : IsSepFor O' O) - (A : Type@{i}), - (forall (x y : A), In@{u a i} O (x = y)) -> In@{u a i} O' A. - Proof. - intros [O'|O'] [O|O]; - [ exact (@Os1.inSepO_from_inO_paths@{u a i iplus} O' O) - | contradiction - | contradiction - | exact (@Os2.inSepO_from_inO_paths@{u a i iplus} O' O) ]. - Defined. - -End ReflectiveSubuniverses_FamUnion. +(** ** Separated subuniverses *) -(** For examples of reflective subuniverses, see the examples of modalities listed in the file [Modality], and also [Localization]. *) +(** For any subuniverse [O], a type is [O]-separated iff all its identity types are [O]-modal. We will study these further in [Separated.v], but we put the definition here because it's needed in [Descent.v]. *) +Definition Sep (O : Subuniverse) : Subuniverse. +Proof. + unshelve econstructor. + - intros A; exact (forall (x y:A), In O (x = y)). + - exact _. + - intros T U ? f ? x y; cbn in *. + refine (inO_equiv_inO' _ (equiv_ap f^-1 x y)^-1). +Defined. + +Global Instance inO_paths_SepO (O : Subuniverse) + {A : Type} {A_inO : In (Sep O) A} (x y : A) + : In O (x = y) + := A_inO x y. diff --git a/theories/Modalities/Separated.v b/theories/Modalities/Separated.v new file mode 100644 index 00000000000..42298cb567f --- /dev/null +++ b/theories/Modalities/Separated.v @@ -0,0 +1,247 @@ +(* -*- mode: coq; mode: visual-line -*- *) +Require Import HoTT.Basics HoTT.Types HoTT.Cubical.DPath. +Require Import EquivalenceVarieties HProp Fibrations Extensions Pullback NullHomotopy Factorization UnivalenceImpliesFunext PathAny. +Require Import ReflectiveSubuniverse Modality Accessible Localization Descent. +Require Import Truncations.Core. +Require Import Homotopy.Suspension. + +Local Open Scope path_scope. +Local Open Scope subuniverse_scope. + +(** * Subuniverses of separated types *) + +(** The basic reference for subuniverses of separated types is + - Christensen, Opie, Rijke, and Scoccola, "Localization in Homotopy Type Theory", https://arxiv.org/abs/1807.04155. +hereinafter referred to as "CORS". *) + +(** ** Definition *) + +(** The definition is in [ReflectiveSubuniverse.v]. *) + +(** ** Basic properties *) + +(** Lemma 2.15 of CORS: If [O] is accessible, so is [Sep O]. Its generators are the suspension of those of [O], in the following sense: *) +Definition susp_localgen (f : LocalGenerators@{a}) : LocalGenerators@{a}. +Proof. + econstructor; intros i. + exact (functor_susp (f i)). +Defined. + +Global Instance isaccrsu_sep (O : Subuniverse) `{IsAccRSU O} + : IsAccRSU (Sep O). +Proof. + exists (susp_localgen (acc_lgen O)). + intros A; split; intros A_inO. + { intros i. + apply (ooextendable_iff_functor_susp (acc_lgen O i)). + intros [x y]. cbn in *. + refine (ooextendable_postcompose' _ _ _ _ _). + 2:apply inO_iff_islocal; exact (A_inO x y). + intros b. + apply equiv_dp_const. } + { intros x y. + apply (inO_iff_islocal O); intros i. + specialize (A_inO i). + refine (ooextendable_postcompose' _ _ _ _ _). + 2:exact (fst (ooextendable_iff_functor_susp (acc_lgen O i) _) A_inO (x,y)). + intros b. + symmetry; apply equiv_dp_const. } +Defined. + +Definition susp_nullgen (S : NullGenerators@{a}) : NullGenerators@{a}. +Proof. + econstructor; intros i. + exact (Susp (S i)). +Defined. + +Global Instance isaccmodality_sep (O : Subuniverse) `{IsAccModality O} + : IsAccModality (Sep O). +Proof. + exists (susp_nullgen (acc_ngen O)). + intros A; split; intros A_inO. + { intros i. + apply (ooextendable_compose _ (functor_susp (fun _:acc_ngen O i => tt)) (fun _:Susp Unit => tt)). + 1:apply ooextendable_equiv, isequiv_contr_contr. + apply (ooextendable_iff_functor_susp (fun _:acc_ngen O i => tt)). + intros [x y]. + refine (ooextendable_postcompose' _ _ _ _ _). + 2:apply inO_iff_isnull; exact (A_inO x y). + intros b. + apply equiv_dp_const. } + { intros x y. + apply (inO_iff_isnull O); intros i. + specialize (A_inO i). + assert (ee : ooExtendableAlong (functor_susp (fun _:acc_ngen O i => tt)) (fun _ => A)). + { refine (cancelL_ooextendable _ _ (fun _ => tt) _ A_inO). + apply ooextendable_equiv. + apply isequiv_contr_contr. } + assert (e := fst (ooextendable_iff_functor_susp (fun _:acc_ngen O i => tt) _) ee (x,y)). + cbn in e. + refine (ooextendable_postcompose' _ _ _ _ e). + intros b. + symmetry; apply equiv_dp_const. } +Defined. + +(** Remark 2.16(1) of CORS *) +Global Instance O_leq_SepO (O : ReflectiveSubuniverse) + : O <= Sep O. +Proof. + intros A ? x y; exact _. +Defined. + +(** Part of Remark 2.16(2) of CORS *) +Definition in_SepO_embedding (O : Subuniverse) + {A B : Type} (i : A -> B) `{IsEmbedding i} `{In (Sep O) B} + : In (Sep O) A. +Proof. + intros x y. + refine (inO_equiv_inO' _ (equiv_ap_isembedding i x y)^-1). +Defined. + +Global Instance in_SepO_hprop (O : ReflectiveSubuniverse) + {A : Type} `{IsHProp A} + : In (Sep O) A. +Proof. + srapply (in_SepO_embedding O (const tt)). + intros x y; exact _. +Defined. + +(** Remark 2.16(4) of CORS *) +Definition sigma_closed_SepO (O : Modality) {A : Type} (B : A -> Type) + `{A_inO : In (Sep O) A} `{B_inO : forall a, In (Sep O) (B a)} + : In (Sep O) (sig B). +Proof. + intros [x u] [y v]. + specialize (A_inO x y). + pose proof (fun p:x=y => B_inO y (p # u) v). + refine (inO_equiv_inO' _ (equiv_path_sigma B _ _)). +Defined. + +(** Lemma 2.17 of CORS *) +Global Instance issurjective_to_SepO (O : ReflectiveSubuniverse) (X : Type) + `{Reflects (Sep O) X} + : IsSurjection (to (Sep O) X). +Proof. + pose (im := himage (to (Sep O) X)). + pose proof (in_SepO_embedding O (factor2 im)). + pose (s := O_rec (factor1 im)). + assert (h : factor2 im o s == idmap). + - apply O_indpaths; intros x; subst s. + rewrite O_rec_beta. + apply fact_factors. + - apply BuildIsSurjection. + intros z. + specialize (h z); cbn in h. + set (w := s z) in *. + destruct w as [w1 w2]. + destruct h. + exact w2. +Defined. + +(** Proposition 2.18 of CORS. *) +Definition almost_inSepO_typeO@{i j} `{Univalence} + (O : ReflectiveSubuniverse) (A B : Type_@{i j} O) + : { Z : Type@{i} & In O Z * (Z <~> (A = B)) }. +Proof. + exists (A <~> B); split. + - exact _. + - refine (equiv_path_TypeO O A B oE _). + apply equiv_path_universe. +Defined. + +(** Lemma 2.21 of CORS *) +Global Instance inSepO_sigma (O : ReflectiveSubuniverse) + {X : Type} {P : X -> Type} `{In (Sep O) X} `{forall x, In O (P x)} + : In (Sep O) (sig P). +Proof. + intros u v. + refine (inO_equiv_inO' _ (equiv_path_sigma P _ _)). +Defined. + +(** Proposition 2.22 of CORS (in funext-free form). *) +Global Instance reflectsD_SepO (O : ReflectiveSubuniverse) + {X : Type} `{Reflects (Sep O) X} + : ReflectsD (Sep O) O X. +Proof. + srapply reflectsD_from_inO_sigma. +Defined. + +(** Once we know that [Sep O] is a reflective subuniverse, this will mean that [O << Sep O]. *) + +(** And now the version with funext. *) +Definition isequiv_toSepO_inO `{Funext} (O : ReflectiveSubuniverse) + {X : Type} `{Reflects (Sep O) X} + (P : O_reflector (Sep O) X -> Type) `{forall x, In O (P x)} + : IsEquiv (fun g : (forall y, P y) => g o to (Sep O) X) + := isequiv_ooextendable _ _ (extendable_to_OO P). + +Definition equiv_toSepO_inO `{Funext} (O : ReflectiveSubuniverse) + {X : Type} `{Reflects (Sep O) X} + (P : O_reflector (Sep O) X -> Type) `{forall x, In O (P x)} + : (forall y, P y) <~> (forall x, P (to (Sep O) X x)) + := Build_Equiv _ _ _ (isequiv_toSepO_inO O P). + +(** TODO: Actually prove this, and put it somewhere more appropriate. *) +Section JoinConstruction. + Universes i j. + Context {X : Type@{i}} {Y : Type@{j}} (f : X -> Y) + (ls : forall (y1 y2 : Y), + @sig@{j j} Type@{i} (fun (Z : Type@{i}) => Equiv@{i j} Z (y1 = y2))). + Definition jc_image@{} : Type@{i}. Admitted. + Definition jc_factor1@{} : X -> jc_image. Admitted. + Definition jc_factor2@{} : jc_image -> Y. Admitted. + Definition jc_factors@{} : jc_factor2 o jc_factor1 == f. Admitted. + Global Instance jc_factor1_issurj@{} : IsSurjection jc_factor1. Admitted. + Global Instance jc_factor2_isemb : IsEmbedding jc_factor2. Admitted. +End JoinConstruction. + +(** We'd like to say that the universe of [O]-modal types is [O]-separated, i.e. belongs to [Sep O]. But since a given subuniverse like [Sep O] lives only on a single universe size, trying to say that in the naive way yields a universe inconsistency. *) +Fail Instance: forall (O : ReflectiveSubuniverse), In (Sep O) (Type_ O). + +(** Instead, we do as in Lemma 2.19 of CORS and prove the morally-equivalent "descent" property, using Lemma 2.18 and the join construction. *) +Global Instance SepO_lex_leq `{Univalence} + (O : ReflectiveSubuniverse) {X : Type} `{Reflects (Sep O) X} + : Descends (Sep O) O X. +Proof. + assert (e : forall (P : X -> Type_ O), + { Q : (O_reflector (Sep O) X -> Type_ O) & + forall x, Q (to (Sep O) X x) <~> P x }). + 2:{ unshelve econstructor; intros P' P_inO; + pose (P := fun x => (P' x; P_inO x) : Type_ O); + pose (ee := e P). + - exact ee.1. + - exact _. + - intros x; cbn; apply ee.2. } + intros P. + assert (ls : forall A B : Type_ O, { Z : Type & Z <~> (A = B) }). + { intros A B. + pose (q := almost_inSepO_typeO O A B). + exact (q.1; snd q.2). } + pose (p := jc_factor2 P ls). + set (J := jc_image P ls) in p. + assert (In (Sep O) J). + { intros x y. + pose (q := almost_inSepO_typeO O (p x) (p y)). + refine (inO_equiv_inO' q.1 _). + refine (_ oE _). + - symmetry; srapply (equiv_ap_isembedding p). + - exact (snd q.2). } + pose (O_rec (O := Sep O) (jc_factor1 P ls)). + exists (p o j). + intros x; subst p j. + rewrite O_rec_beta. + apply equiv_path. + exact ((jc_factors P ls x)..1). +Defined. + +(** Once we know that [Sep O] is a reflective subuniverse, this will imply [O <<< Sep O], and that if [Sep O] is accessible (such as if [O] is) then [Type_ O] belongs to its accessible lifting (see [inO_TypeO_lex_leq]. *) + + +(** ** Reflectiveness of [Sep O] *) + +(** TODO *) + + +(** ** Left-exactness properties *) + +(** Nearly all of these are true in the generality of a pair of reflective subuniverses with [O <<< O'] and/or [O' <= Sep O], and as such can be found in [Descent.v]. *) diff --git a/theories/Modalities/Topological.v b/theories/Modalities/Topological.v index b44e14aeb1f..f3b41510ab2 100644 --- a/theories/Modalities/Topological.v +++ b/theories/Modalities/Topological.v @@ -11,25 +11,19 @@ Local Open Scope path_scope. (** A topological localization -- or, as we will say, a topological nullification -- is a nullification at a family of hprops, or more generally an accessible modality whose generators of accessibility are all hprops. This is not quite the same as Lurie's definition: in Higher Topos Theory, a topological localization is an accessible *left exact* localization at a pullback-stable class generated by a set of monomorphisms. "Pullback-stable class generated by" is roughly incorporated into our internal notion of accessibility, so the main new difference here is that when the generation is internal in this way, the localization at a family of hprops is *automatically* left exact. *) -Module Topological_Modalities_Theory - (Os : Modalities) (Acc : Accessible_Modalities Os). +Notation Topological O := (forall i, IsHProp (acc_ngen O i)). - Module Export Acc_Lex_Theory := Accessible_Lex_Modalities_Theory Os Acc. +(** ** Topological modalities are lex *) - Notation Topological O := (forall i, IsHProp (acc_gen O i)). +(** We prove left-exactness by proving that the universe of modal types is modal, using univalence. It's unclear whether univalence is necessary or not in general; in one special case (open modalities) funext suffices. But it's plausible that it would be necessary in general, because lex-ness of nullification is a statement about the path-spaces of a HIT, and characterizing those in any way usually requires some amount of univalence. *) - (** ** Topological modalities are lex *) - - (** We prove left-exactness by proving that the universe of modal types is modal. Of course, this requires univalence. *) - - Global Instance lex_topological `{Univalence} - (O : Modality) `{Topological O} +Global Instance lex_topological `{Univalence} + (O : Modality) `{IsAccModality O} `{Topological O} : Lex O. - Proof. - apply lex_inO_typeO. - refine (snd (inO_iff_isnull O _) _); intros i. - refine (equiv_inverse (equiv_ooextendable_isequiv _ _) _). - simple refine (isequiv_adjointify _ _ _ _); simpl. +Proof. + snrapply lex_from_inO_typeO; [ exact _ | intros i ]. + apply ((equiv_ooextendable_isequiv _ _)^-1%equiv). + srapply isequiv_adjointify; cbn. - intros B _. refine ((forall a, B a) ; _). exact _. @@ -53,31 +47,14 @@ Module Topological_Modalities_Theory unfold composeD in e; simpl in e. refine (_ oE (Build_Equiv _ _ _ e)^-1). exact (equiv_contr_forall _). - Defined. - -End Topological_Modalities_Theory. - -(** In particular, a nullification at a family of hprops is topological and therefore lex. *) - -Module Import Topological_Nullification_Theory := - Topological_Modalities_Theory - Nullification_Modalities - Accessible_Nullification. - -(** It isn't necessary to declare these as global instances, since typeclass search can find them automatically. But we want to state them explicitly here for exposition, so we make them local instances. *) -Local Instance topological_nullification - (O : Nullification_Modality) `{forall i, IsHProp (unNul O i)} -: Topological O. -Proof. - exact _. Defined. -(** Note the hypothesis of [Univalence] required for this one. It's unclear whether this is necessary or not in general; in one special case (open modalities) funext suffices. But it's plausible that it would be necessary in general, because lex-ness of nullification is a statement about the path-spaces of a HIT, and characterizing those in any way usually requires some amount of univalence. *) -Local Instance lex_nullification `{Univalence} - (O : Nullification_Modality) `{forall i, IsHProp (unNul O i)} -: Lex O. + +Global Instance lex_nullification `{Univalence} + (S : NullGenerators) `{forall i, IsHProp (S i)} +: Lex (Nul S). Proof. - exact _. + rapply lex_topological. Defined. @@ -87,31 +64,25 @@ Defined. (** This is kind of annoying to prove, not just because the proof is fiddly, but because we have to pass back and forth between different generating families for the "same" modality. It's a bit easier to prove it about nullifications than about arbitrary accessible lex modalities. *) -Module Topological_Lex. - Import NulM. - Import AccNulM. - Module Import LexNulM := Lex_Modalities_Theory Nullification_Modalities. - - Definition topological_lex_trunc_acc `{Funext} +Definition topological_lex_trunc_acc `{Funext} (B : NullGenerators) {Olex : Lex (Nul B)} (n : trunc_index) (gtr : forall a, IsTrunc n (ngen_type B a)) : { D : NullGenerators & (forall c, IsHProp (ngen_type D c)) * - OeqO (Nul D) (Nul B) }. - Proof. + (Nul D <=> Nul B) }. +Proof. destruct n. { exists (Build_NullGenerators Empty (fun _ => Unit)). - split; [ exact _ | intros X ]. - split; intros _; [ | intros [] ]. - apply inO_iff_isnull; intros i. + split; [ exact _ | split; intros X _; [ | intros [] ] ]. + intros i. apply ooextendable_equiv, isequiv_contr_contr. } pose (O := Nul B). - pose (OeqB := reflexivity O : OeqO O (Nul B)). - fold O in Olex. + pose (OeqB := reflexive_O_eq O : O <=> (Nul B)). + change (Nul B) with O in Olex. clearbody O OeqB. revert B OeqB gtr. induction n; intros B OeqB gtr. - { exists B; split; [ assumption | intros ?; reflexivity ]. } + { exists B; split; [ assumption | reflexivity ]. } pose (A := ngen_indices B). pose (C := A + { a:A & B(a) * B(a) }). pose (D := Build_NullGenerators @@ -123,14 +94,13 @@ Module Topological_Lex. { intros [a | [a [b1 b2]]]; [ cbn | exact _ ]. (* Because [trunc_hprop] can't be used as an idmap... *) destruct n; exact _. } - assert (OeqD : OeqO O (Nul D)). - { intros X; split. + assert (OeqD : O <=> (Nul D)). + { split; intros X. - intros X_inO c. assert (Bc : forall a:A, IsConnected O (B a)). { intros a. - apply (@isconnected_OeqO (Nul B) O). - - symmetry; exact _. - - exact (isconnected_acc_gen (Nul B) a). } + rapply (@isconnected_O_leq O (Nul B)). + exact (isconnected_acc_ngen (Nul B) a). } apply (ooextendable_const_isconnected_inO O); [ destruct c as [a | [a [b1 b2]]] | exact X_inO ]. + apply isconnected_from_elim_to_O. @@ -140,8 +110,8 @@ Module Topological_Lex. exists x; intros y; cbn in y. strip_truncations. exact (h y). - + exact _. - - intros Dnull; apply (@inO_OeqO _ _ OeqB). + + cbn. rapply isconnected_paths. + - intros Dnull; rapply (@inO_leq (Nul B) O). intros a; cbn in a; cbn. apply ((equiv_ooextendable_isequiv (unit_name X) (fun _:B a => tt))^-1). @@ -151,13 +121,13 @@ Module Topological_Lex. intros x; unfold composeD; cbn. apply equiv_path_arrow. } refine ((isconnected_elim (Nul D) (A := D (inl a)) _ _).1). - { apply isconnected_acc_gen. } + { rapply isconnected_acc_ngen. } intros b; cbn in b. strip_truncations. assert (bc : IsConnMap (Nul D) (unit_name b)). { intros x; unfold hfiber. apply (isconnected_equiv (Nul D) (b = x) (equiv_contr_sigma _)^-1). - apply (isconnected_acc_gen (Nul D) (inr (a;(b,x)))). } + rapply (isconnected_acc_ngen (Nul D) (inr (a;(b,x)))). } pose (p := conn_map_elim (Nul D) (unit_name b) (fun u => f b = f u) (fun _ => 1)). exists (f b ; p); intros [x q]. @@ -173,6 +143,4 @@ Module Topological_Lex. refine (transitivity EeqD _). refine (transitivity _ OeqB). symmetry; assumption. - Defined. - -End Topological_Lex. +Defined. diff --git a/theories/Pointed/Loops.v b/theories/Pointed/Loops.v index e4861781e18..7acdd4b2e75 100644 --- a/theories/Pointed/Loops.v +++ b/theories/Pointed/Loops.v @@ -3,8 +3,6 @@ Require Import HSet Fibrations Factorization HoTT.Truncations HProp. Require Import UnivalenceImpliesFunext. Require Import Pointed.Core Pointed.pMap Pointed.pEquiv Pointed.pHomotopy. -Import TrM. - Local Open Scope pointed_scope. Local Open Scope path_scope. diff --git a/theories/Pointed/pSusp.v b/theories/Pointed/pSusp.v index 25e1d36c62c..103b52f206f 100644 --- a/theories/Pointed/pSusp.v +++ b/theories/Pointed/pSusp.v @@ -9,7 +9,6 @@ Require Import Pointed.pEquiv. Require Import Homotopy.Suspension. Require Import Homotopy.Freudenthal. Require Import Truncations. -Import TrM. Generalizable Variables X A B f g n. diff --git a/theories/Pointed/pTrunc.v b/theories/Pointed/pTrunc.v index 460a5b2dcc4..db5eac29e92 100644 --- a/theories/Pointed/pTrunc.v +++ b/theories/Pointed/pTrunc.v @@ -5,7 +5,6 @@ Require Import Pointed.pMap. Require Import Pointed.pEquiv. Require Import Pointed.Loops. Require Import Truncations. -Import TrM. Local Open Scope pointed_scope. @@ -113,4 +112,3 @@ Proof. { apply ptr_functor, e. } exact _. Defined. - diff --git a/theories/Spaces/BAut.v b/theories/Spaces/BAut.v index 03166c4edf4..e0cb464d5bd 100644 --- a/theories/Spaces/BAut.v +++ b/theories/Spaces/BAut.v @@ -3,7 +3,6 @@ Require Import HoTT.Basics HoTT.Types HProp. Require Import Constant Factorization UnivalenceImpliesFunext. Require Import Modalities.Modality HoTT.Truncations. Require Export Algebra.ooGroup Algebra.Aut. -Import TrM. Local Open Scope path_scope. diff --git a/theories/Spaces/BAut/Bool.v b/theories/Spaces/BAut/Bool.v index c7de3d47b2c..b6bdd02717c 100644 --- a/theories/Spaces/BAut/Bool.v +++ b/theories/Spaces/BAut/Bool.v @@ -2,7 +2,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import Constant Factorization UnivalenceImpliesFunext. Require Import Modalities.Modality HoTT.Truncations. -Import TrM. Require Import Spaces.BAut. Local Open Scope trunc_scope. diff --git a/theories/Spaces/BAut/Rigid.v b/theories/Spaces/BAut/Rigid.v index 7731c56ab6b..eba73c5e524 100644 --- a/theories/Spaces/BAut/Rigid.v +++ b/theories/Spaces/BAut/Rigid.v @@ -2,7 +2,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import HProp UnivalenceImpliesFunext Fibrations. Require Import Modalities.Modality Truncations. -Import TrM. Require Import Spaces.BAut. Local Open Scope trunc_scope. @@ -29,7 +28,7 @@ Global Instance contr_baut_rigid `{Univalence} {A : Type} `{IsRigid A} Proof. refine (contr_trunc_conn (Tr 0)). intros Z W; baut_reduce. - refine (trunc_equiv (A <~> A) + refine (trunc_equiv (n := -1) (A <~> A) (path_baut (point (BAut A)) (point (BAut A)))). Defined. diff --git a/theories/Spaces/Finite.v b/theories/Spaces/Finite.v index f4f95fac2c4..05f37b20d13 100644 --- a/theories/Spaces/Finite.v +++ b/theories/Spaces/Finite.v @@ -11,7 +11,6 @@ Require Import EquivalenceVarieties. Require Import UnivalenceImpliesFunext. Require Import Truncations. Require Import Colimits.Quotient. -Import TrM. Local Open Scope path_scope. Local Open Scope nat_scope. @@ -798,7 +797,7 @@ Proof. refine (decidable_equiv _ (hfiber_fibration x P)^-1 _). refine (detachable_image_finite pr1 x). - assumption. (** Why doesn't Coq find this? *) - - apply mapinO_pr1; exact _. (** Why doesn't Coq find this? *) + - apply (mapinO_pr1 (Tr (-1))). (** Why doesn't Coq find this? *) Defined. (** ** Quotients *) @@ -910,12 +909,12 @@ Definition leq_inj_finite `{Funext} {X Y} {fX : Finite X} {fY : Finite Y} (f : X -> Y) (i : IsEmbedding f) : fcard X <= fcard Y. Proof. - assert (MapIn (-1)%trunc f) by exact _. clear i. + assert (MapIn (Tr (-1)) f) by exact _. clear i. destruct fX as [n e]; simpl. destruct fY as [m e']; simpl. strip_truncations. pose (g := e' o f o e^-1). - assert (MapIn (-1)%trunc g) by (unfold g; exact _). + assert (MapIn (Tr (-1)) g) by (unfold g; exact _). clearbody g. clear e e'. generalize dependent m. induction n as [|n IHn]. { intros; exact tt. } @@ -924,7 +923,7 @@ Proof. destruct m as [|m]. { elim (g (inr tt)). } pose (h := (fin_transpose_last_with m (g (inr tt)))^-1 o g). - assert (MapIn (-1)%trunc h) by (unfold h; exact _). + assert (MapIn (Tr (-1)) h) by (unfold h; exact _). assert (Ha : forall a:Fin n, is_inl (h (inl a))). { intros a. remember (g (inl a)) as b eqn:p. @@ -951,7 +950,7 @@ Proof. apply fin_transpose_last_with_last. } rewrite q; exact tt. } exact (IHn m (unfunctor_sum_l h Ha) - (mapinO_unfunctor_sum_l (-1)%trunc h Ha Hb)). + (mapinO_unfunctor_sum_l (Tr (-1)) h Ha Hb)). Qed. (** ** Initial segments of [nat] *) diff --git a/theories/Spaces/No/Core.v b/theories/Spaces/No/Core.v index fb85d4e30e2..fdda69778b0 100644 --- a/theories/Spaces/No/Core.v +++ b/theories/Spaces/No/Core.v @@ -2,7 +2,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import UnivalenceImpliesFunext TruncType HSet. Require Import HoTT.Truncations. -Import TrM. Local Open Scope nat_scope. Local Open Scope path_scope. diff --git a/theories/Spaces/Torus/TorusHomotopy.v b/theories/Spaces/Torus/TorusHomotopy.v index edf933e3176..0eff33a6bed 100644 --- a/theories/Spaces/Torus/TorusHomotopy.v +++ b/theories/Spaces/Torus/TorusHomotopy.v @@ -10,7 +10,6 @@ Require Import HIT.Circle. Require Import Truncations. Require Import Homotopy.HomotopyGroup. Require Import UnivalenceImpliesFunext. -Import TrM. Require Import Torus. Require Import TorusEquivCircles. diff --git a/theories/Spaces/Universe.v b/theories/Spaces/Universe.v index 89629b7ca71..673d8dd45b5 100644 --- a/theories/Spaces/Universe.v +++ b/theories/Spaces/Universe.v @@ -2,7 +2,6 @@ Require Import HoTT.Basics HoTT.Types. Require Import HProp UnivalenceImpliesFunext Fibrations. Require Import Modalities.Modality HoTT.Truncations. -Import TrM. Require Import Spaces.BAut Spaces.BAut.Rigid. Require Import ExcludedMiddle. diff --git a/theories/Spectra/Spectrum.v b/theories/Spectra/Spectrum.v index 297c281c57a..babf6ded115 100644 --- a/theories/Spectra/Spectrum.v +++ b/theories/Spectra/Spectrum.v @@ -7,8 +7,6 @@ Require Import HoTT.Tactics. Require Import Pointed. Require Import HoTT.Truncations. -Import TrM. - Local Open Scope nat_scope. Local Open Scope path_scope. Local Open Scope equiv_scope. diff --git a/theories/Tests.v b/theories/Tests.v index 31f1081c2cc..4162097447c 100644 --- a/theories/Tests.v +++ b/theories/Tests.v @@ -24,17 +24,6 @@ Check (@ispointed_susp Set). (** Check that nested sigma-type notation didn't get clobbered by surreal cuts *) Check ({ l : Unit & { n : Unit & Unit }}). -(** Regression check issue #744 *) -Module Foo (Os : ReflectiveSubuniverses). - Module Import Os_Theory := ReflectiveSubuniverses_Theory Os. - Goal Unit. - let lem' := preconcat_any @to_O_natural_compose in - pose proof lem' as H. - let test := left_associate_concat_in H in - pose test. - Admitted. -End Foo. - (** Test 1 from issue #754 *) Module Issue754_1. Inductive nat : Type1 := diff --git a/theories/Truncations/Connectedness.v b/theories/Truncations/Connectedness.v index 86b9a486c22..62eb43fcdab 100644 --- a/theories/Truncations/Connectedness.v +++ b/theories/Truncations/Connectedness.v @@ -9,9 +9,9 @@ Require Import HProp. Require Import EquivalenceVarieties. Require Import Extensions. Require Import Factorization. -Require Export Modality. (* [Export] since the actual definitions of connectednes appear there, in the generality of a modality. *) +Require Export Modalities.Modality. (* [Export] since the actual definitions of connectednes appear there, in the generality of a modality. *) +Require Import Modalities.Descent Modalities.Separated. Require Import Truncations.Core. -Import TrM. Local Open Scope path_scope. Local Open Scope trunc_scope. @@ -63,30 +63,24 @@ Defined. (** We can't make both of these [Instance]s, as that would result in infinite loops. *) Global Instance conn_pointed_type {n : trunc_index} {A : Type} (a0:A) - `{IsConnMap n _ _ (unit_name a0)} : IsConnected n.+1 A | 1000. + `{IsConnMap n _ _ (unit_name a0)} + : IsConnected n.+1 A | 1000. Proof. - apply isconnected_from_elim. - intros C HC f. exists (f a0). - refine (conn_map_elim n (unit_name a0) _ (fun _ => idpath)). + apply isconnected_conn_map_to_unit. + rapply (OO_cancelR_conn_map (Tr n.+1) (Tr n) (unit_name a0) (fun _:A => tt)). Defined. Definition conn_point_incl {n : trunc_index} {A : Type} (a0:A) - `{IsConnected n.+1 A} : IsConnMap n (unit_name a0). + `{IsConnected n.+1 A} + : IsConnMap n (unit_name a0). Proof. - apply conn_map_from_extension_elim. - intros P ?. set (PP := fun a => BuildTruncType n (P a)). - assert (QQ := isconnected_elim n.+1 (TruncType n) PP). - destruct QQ as [[Q0 HQ] e]. - assert (e' := fun a => ap trunctype_type (e a)); simpl in e'. clear HQ e. - intros d. set (d0 := d tt). - exists (fun a => (transport idmap (e' a0 @ (e' a)^) d0)). - intros []. change (d tt) with (transport idmap 1 d0). - apply ap10, ap, concat_pV. + rapply (OO_cancelL_conn_map (Tr n.+1) (Tr n) (unit_name a0) (fun _:A => tt)). + apply O_lex_leq_Tr. Defined. Hint Immediate conn_point_incl : typeclass_instances. -(** TODO: generalise the above to any map with a section. *) +(** Note that [OO_cancelR_conn_map] and [OO_cancelL_conn_map] (Proposition 2.31 of CORS) generalize the above statements to 2/3 of a 2-out-of-3 property for connected maps, for any reflective subuniverse and its subuniverse of separated types. If useful, we could specialize that more general form explicitly to truncations. *) End Extensions. @@ -130,7 +124,7 @@ Defined. (** ** 0-connectedness *) -(** To be 0-connected is the same as to be (-1)-connected and that any two points are merely equal. *) +(** To be 0-connected is the same as to be (-1)-connected and that any two points are merely equal. TODO: This should also be generalized to separated subuniverses (CORS Remark 2.35). *) Definition merely_path_is0connected `{Univalence} (A : Type) `{IsConnected 0 A} (x y : A) : merely (x = y). @@ -177,7 +171,7 @@ Proof. Defined. (* Truncation preserves connectedness. Note that this is for different levels. *) -Global Instance isconnected_trunc {X : Type} n m `{IsConnected n X} +Global Instance isconnected_trunc {X : Type} (n m : trunc_index) `{IsConnected n X} : IsConnected n (Tr m X). Proof. unfold IsConnected. diff --git a/theories/Truncations/Core.v b/theories/Truncations/Core.v index ad0c6cce26b..33605dad48e 100644 --- a/theories/Truncations/Core.v +++ b/theories/Truncations/Core.v @@ -2,7 +2,7 @@ Require Import Basics Types. Require Import TruncType HProp UnivalenceImpliesFunext. -Require Import Modalities.Modality Modalities.Identity. +Require Import Modalities.Modality Modalities.Identity Modalities.Descent. (** * Truncations of types, in all dimensions. *) @@ -26,7 +26,7 @@ However, while we are faking our higher-inductives anyway, we can take some shor Module Export Trunc. Delimit Scope trunc_scope with trunc. -Private Inductive Trunc (n : trunc_index) (A :Type) : Type := +Cumulative Private Inductive Trunc (n : trunc_index) (A :Type) : Type := tr : A -> Trunc n A. Bind Scope trunc_scope with Trunc. Arguments tr {n A} a. @@ -49,104 +49,31 @@ Definition Trunc_rec {n A X} `{IsTrunc n X} : (A -> X) -> (Trunc n A -> X) := Trunc_ind (fun _ => X). -(** [Trunc] is a modality *) - -(** We will define a truncation modality to be parametrized by a [trunc_index]. However, as described in Modality.v, we don't want to simply define [Truncation_Modalities.Modality] to *be* [trunc_index]; we want to think of the truncation modality as being *derived from* rather than *identical to* its truncation index. In particular, there is a coercion [O_reflector] from [Modality] to [Funclass], but we don't want Coq to print things like [2 X] to mean [Trunc 2 X]. However, in the special case of truncation, it is nevertheless convenient for [Truncation_Modalities.Modality] to be definitionally equal to [trunc_index], so that we can call modality functions (particularly connectedness functions) passing a truncation index directly. - -These may seem like contradictory requirements, but it appears to be possible to satisfy them because coercions don't unfold definitions. Thus, rather than a record wrapper, we define a *definitional* wrapper [Truncation_Modality] around [trunc_index], and a notation [Tr] for the identity. We will define [Truncation_Modalities.Modality] to be [Truncation_Modality] and declare the identity as a coercion; thus a [Truncation_Modality] can be used as a modality and therefore also as a function (via the [O_reflector] coercion). However, the identity from [trunc_index] to [Truncation_Modality] is not a coercion, so we don't get notation like [2 X]. *) -Definition Truncation_Modality := trunc_index. -Definition Tr : trunc_index -> Truncation_Modality := idmap. - -Module Truncation_Modalities <: Modalities. - - Definition Modality : Type2@{u a} := Truncation_Modality. - - Definition O_reflector (n : Modality@{u u'}) A := Trunc n A. - - Definition In (n : Modality@{u u'}) A := IsTrunc n A. - - Definition O_inO (n : Modality@{u u'}) A : In n (O_reflector n A). - Proof. - unfold In, O_reflector; exact _. - Defined. - - Definition to (n : Modality@{u u'}) A := @tr n A. - - Definition inO_equiv_inO (n : Modality@{u u'}) - (A : Type@{i}) (B : Type@{j}) Atr f feq - : let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - In n B - := @trunc_equiv A B f n Atr feq. - - Definition hprop_inO `{Funext} (n : Modality@{u u'}) A - : IsHProp (In n A). - Proof. - unfold In; exact _. - Defined. - - Definition O_ind_internal - : forall (n : Modality@{u a}) - (A : Type@{i}) (B : O_reflector n A -> Type@{j}) - (B_inO : forall oa, In@{u a j} n (B oa)), - let gei := ((fun x => x) : Type@{i} -> Type@{k}) in - let gej := ((fun x => x) : Type@{j} -> Type@{k}) in - (forall a, B (to n A a)) -> forall a, B a - := @Trunc_ind. - - Definition O_ind_beta_internal (n : Modality@{u u'}) - A B Btr f a - : O_ind_internal n A B Btr f (to n A a) = f a - := 1. - - Definition minO_paths (n : Modality@{u a}) - (A : Type@{i}) (Atr : In@{u a i} n A) (a a' : A) - : In@{u a i} n (a = a'). - Proof. - unfold In in *; exact _. - Defined. - - Definition IsSepFor@{u a} (n' n : Modality@{u a}) : Type@{u} - := (paths@{a} n.+1 n'). - - Definition inO_paths_from_inSepO@{u a i iplus} - (n' n : Modality@{u a}) (sep : IsSepFor n' n) - (A : Type@{i}) (A_inO : In@{u a i} n' A) (x y : A) - : In@{u a i} n (x = y). - Proof. - destruct sep; unfold In in *. - exact _. - Defined. - - Definition inSepO_from_inO_paths@{u a i iplus} - (n' n : Modality@{u a}) (sep : IsSepFor n' n) - (A : Type@{i}) (inO_pathsA : forall (x y : A), In@{u a i} n (x = y)) - : In@{u a i} n' A. - Proof. - destruct sep; unfold In in *. - intros x y; apply inO_pathsA. - Defined. - -End Truncation_Modalities. - -(** If you import the following module [TrM], then you can call all the modality functions with a [trunc_index] as the modality parameter, since we defined [Truncation_Modalities.Modality] to be [trunc_index]. *) -Module Import TrM := Modalities_Theory Truncation_Modalities. -(** If you don't import it, then you'll need to write [TrM.function_name] or [TrM.RSU.function_name] depending on whether [function_name] pertains only to modalities or also to reflective subuniverses. (Having to know this is a bit unfortunate, but apparently the fact that [TrM] [Export]s reflective subuniverses still doesn't make the fields of the latter accessible as [TrM.field].) *) -Export TrM.Coercions. -Export TrM.RSU.Coercions. - -(** Here is the additional coercion promised above. *) -Coercion Truncation_Modality_to_Modality := idmap : Truncation_Modality -> Modality. - -(** Sometimes we also need an explicit definition of the truncation reflective subuniverses. *) -Module Truncation_RSUs := Modalities_to_ReflectiveSubuniverses Truncation_Modalities. +(** ** [Trunc] is a modality *) + +Definition Tr (n : trunc_index) : Modality. +Proof. + srapply (Build_Modality (IsTrunc n)). + - intros A B ? f ?; apply (trunc_equiv A f). + - exact (Trunc n). + - intros; apply istrunc_truncation. + - intros A; apply tr. + - intros A B ? f oa; cbn in *. + exact (Trunc_ind B f oa). + - intros; reflexivity. +Defined. + +(** We don't usually declare modalities as coercions, but this particular one is convenient so that lemmas about (for instance) connected maps can be applied to truncation modalities without the user/reader needing to be (particularly) aware of the general notion of modality. *) +Coercion Tr : trunc_index >-> Modality. +(** However, if the coercion is not printed, then we get things like [Tr (-1) X] being printed as [(-1) X], which is terribly confusing. So we tell Coq to always print this coercion. This does mean that although the user can type things like [IsConnected n X], it will always be displayed back as [IsConnected (Tr n) X]. *) +Add Printing Coercion Tr. Section TruncationModality. Context (n : trunc_index). Definition trunc_iff_isequiv_truncation (A : Type) : IsTrunc n A <-> IsEquiv (@tr n A) - := inO_iff_isequiv_to_O n A. + := inO_iff_isequiv_to_O (Tr n) A. Global Instance isequiv_tr A `{IsTrunc n A} : IsEquiv (@tr n A) := fst (trunc_iff_isequiv_truncation A) _. @@ -161,39 +88,35 @@ Section TruncationModality. (** ** Functoriality *) - Definition Trunc_functor {X Y : Type} (f : X -> Y) - : Tr n X -> Tr n Y - := O_functor n f. + (** Since a modality lives on a single universe, by default if we simply define [Trunc_functor] to be [O_functor] then it would force [X] and [Y] to live in the same universe. But since we defined [Trunc] as a cumulative inductive, if we add universe annotations we can make [Trunc_functor] more universe-polymorphic than [O_functor] is. This is sometimes useful. *) + Definition Trunc_functor@{i j k} {X : Type@{i}} {Y : Type@{j}} (f : X -> Y) + : Tr@{i} n X -> Tr@{j} n Y + := O_functor@{k k k} (Tr n) f. Global Instance Trunc_functor_isequiv {X Y : Type} (f : X -> Y) `{IsEquiv _ _ f} : IsEquiv (Trunc_functor f) - := isequiv_O_functor n f. + := isequiv_O_functor (Tr n) f. Definition Trunc_functor_equiv {X Y : Type} (f : X <~> Y) : Tr n X <~> Tr n Y - := equiv_O_functor n f. + := equiv_O_functor (Tr n) f. Definition Trunc_functor_compose {X Y Z} (f : X -> Y) (g : Y -> Z) : Trunc_functor (g o f) == Trunc_functor g o Trunc_functor f - := O_functor_compose n f g. + := O_functor_compose (Tr n) f g. Definition Trunc_functor_idmap (X : Type) : @Trunc_functor X X idmap == idmap - := O_functor_idmap n X. + := O_functor_idmap (Tr n) X. Definition isequiv_Trunc_functor {X Y} (f : X -> Y) `{IsEquiv _ _ f} : IsEquiv (Trunc_functor f) - := isequiv_O_functor n f. + := isequiv_O_functor (Tr n) f. Definition equiv_Trunc_prod_cmp `{Funext} {X Y} : Tr n (X * Y) <~> Tr n X * Tr n Y - := equiv_O_prod_cmp n X Y. - - (** Separatedness *) - - Global Instance issepfor_trunc n : IsSepFor n.+1 n - := idpath. + := equiv_O_prod_cmp (Tr n) X Y. End TruncationModality. @@ -233,10 +156,8 @@ Defined. Hint Immediate istruncmap_mapinO_tr : typeclass_instances. - (** It's sometimes convenient to use "infinity" to refer to the identity modality in a similar way. This clashes with some uses in higher topos theory, where "oo-truncated" means instead "hypercomplete", but this has not yet been a big problem. *) Notation oo := purely. -(** Unfortunately, we can't import two or more copies of [Modalities_Theory] at the same time (the most recently imported shadows the other(s)). If we ever want to talk about truncation and include [oo], we may want to define a "union" instance of [Modality]. *) (** ** A few special things about the (-1)-truncation. *) @@ -244,7 +165,7 @@ Local Open Scope trunc_scope. (** We define [merely A] to be an inhabitant of the universe [hProp] of hprops, rather than a type. We can always treat it as a type because there is a coercion, but this means that if we need an element of [hProp] then we don't need a separate name for it. *) -Definition merely (A : Type@{i}) : hProp@{i} := BuildhProp (Trunc (-1) A). +Definition merely (A : Type@{i}) : hProp@{i} := BuildhProp (Tr (-1) A). (** Note that we define [merely] using [Trunc -1] rather than [Tr -1]. These are of course judgmentally equal, but our choice introduces fewer universe parameters, resulting in faster compilation times. The other choice might in theory give Coq an easier time applying general modality theorems to [merely], but currently things seem to be transparent enough that it doesn't matter. *) @@ -254,16 +175,16 @@ Definition hor (P Q : Type) : hProp := merely (P + Q). Notation "A \/ B" := (hor A B) : type_scope. -Definition himage {X Y} (f : X -> Y) := image (-1) f. +Definition himage {X Y} (f : X -> Y) := image (Tr (-1)) f. Definition contr_inhab_prop {A} `{IsHProp A} (ma : merely A) : Contr A. Proof. - refine (@contr_trunc_conn (-1) A _ _); try assumption. + refine (@contr_trunc_conn (Tr (-1)) A _ _); try assumption. refine (contr_inhabited_hprop _ ma). Defined. (** Surjections are the (-1)-connected maps, but they can be characterized more simply since an inhabited hprop is automatically contractible. *) -Notation IsSurjection := (IsConnMap (-1)). +Notation IsSurjection := (IsConnMap (Tr (-1))). Definition BuildIsSurjection {A B} (f : A -> B) : (forall b, merely (hfiber f b)) -> IsSurjection f. @@ -276,7 +197,7 @@ Definition isequiv_surj_emb {A B} (f : A -> B) `{IsSurjection f} `{IsEmbedding f} : IsEquiv f. Proof. - apply (@isequiv_conn_ino_map (-1)); assumption. + apply (@isequiv_conn_ino_map (Tr (-1))); assumption. Defined. (** ** Tactic to remove truncations in hypotheses if possible. *) @@ -291,18 +212,6 @@ Ltac strip_truncations := intro T end. -(** ** Path-spaces of truncations *) - -(** These are just special cases of results about subuniverses of separated types. *) - -Definition path_Tr {n A} (x y : A) - : Tr n (x = y) -> (tr x = tr y :> Tr n.+1 A) - := path_SepO x y. - -Definition equiv_path_Tr `{Univalence} {n A} (x y : A) - : Tr n (x = y) <~> (tr x = tr y :> Tr n.+1 A) - := equiv_path_SepO x y. - (** ** Iterated truncations *) Definition Trunc_min n m X : Tr n (Tr m X) <~> Tr (trunc_index_min n m) X. @@ -346,4 +255,52 @@ Proof. apply Trunc_min. Defined. +(** ** Separatedness and path-spaces of truncations *) + +Section SeparatedTrunc. + +Local Open Scope subuniverse_scope. + +(** The [n.+1]-truncation modality consists of the separated types for the [n]-truncation modality. *) +Global Instance O_eq_Tr (n : trunc_index) + : Tr n.+1 <=> Sep (Tr n). +Proof. + split; intros A A_inO. + - intros x y; exact _. + - exact _. +Defined. + +(** It follows that [Tr n <<< Tr n.+1]. However, it is easier to prove this directly than to go through separatedness. *) +Global Instance O_leq_Tr (n : trunc_index) + : Tr n <= Tr n.+1. +Proof. + intros A ?; exact _. +Defined. + +Global Instance O_strong_leq_Tr (n : trunc_index) + : Tr n << Tr n.+1. +Proof. + srapply O_strong_leq_trans_l. +Defined. + +(** For some reason, this causes typeclass search to spin. *) +Local Instance O_lex_leq_Tr `{Univalence} (n : trunc_index) + : Tr n <<< Tr n.+1. +Proof. + intros A; unshelve econstructor; intros P' P_inO; pose (P := fun x => BuildTruncType n (P' x)). + - refine (Trunc_rec P). + - intros; exact _. + - intros; cbn. reflexivity. +Defined. + +Definition path_Tr {n A} (x y : A) + : Tr n (x = y) -> (tr x = tr y :> Tr n.+1 A) + := path_OO (Tr n.+1) (Tr n) x y. + +Definition equiv_path_Tr `{Univalence} {n A} (x y : A) + : Tr n (x = y) <~> (tr x = tr y :> Tr n.+1 A) + := equiv_path_OO (Tr n.+1) (Tr n) x y. + +End SeparatedTrunc. + (** If you are looking for a theorem about truncation, you may want to read the note "Finding Theorems" in "STYLE.md". *)