|
| 1 | +(** Summary of the formalization (kernel) |
| 2 | +
|
| 3 | +Tip: The Coq command [About ident] prints where the ident was defined |
| 4 | + *) |
| 5 | + |
| 6 | + |
| 7 | +Require Import UniMath.Foundations.PartD. |
| 8 | + |
| 9 | +Require Import UniMath.CategoryTheory.Monads.Monads. |
| 10 | +Require Import UniMath.CategoryTheory.Monads.LModules. |
| 11 | +Require Import UniMath.CategoryTheory.SetValuedFunctors. |
| 12 | +Require Import UniMath.CategoryTheory.HorizontalComposition. |
| 13 | +Require Import UniMath.CategoryTheory.functor_categories. |
| 14 | +Require Import UniMath.CategoryTheory.categories.category_hset. |
| 15 | +Require Import UniMath.CategoryTheory.categories.category_hset_structures. |
| 16 | + |
| 17 | +Require Import UniMath.CategoryTheory.Categories. |
| 18 | +Require Import UniMath.Foundations.Sets. |
| 19 | +Require Import UniMath.CategoryTheory.Epis. |
| 20 | +Require Import UniMath.CategoryTheory.EpiFacts. |
| 21 | + |
| 22 | +Require Import Modules.Prelims.lib. |
| 23 | +Require Import Modules.Prelims.quotientmonad. |
| 24 | +Require Import Modules.Prelims.quotientmonadslice. |
| 25 | +Require Import Modules.Signatures.Signature. |
| 26 | +Require Import Modules.SoftEquations.ModelCat. |
| 27 | +Require Import Modules.Prelims.modules. |
| 28 | + |
| 29 | +Require Import Modules.SoftEquations.quotientrepslice. |
| 30 | +Require Import Modules.SoftEquations.SignatureOver. |
| 31 | +Require Import Modules.SoftEquations.Equation. |
| 32 | +Require Import Modules.SoftEquations.quotientrepslice. |
| 33 | +Require Import Modules.SoftEquations.quotientequation. |
| 34 | + |
| 35 | +Require Import UniMath.CategoryTheory.limits.initial. |
| 36 | +Require Import Modules.SoftEquations.InitialEquationRep. |
| 37 | + |
| 38 | +Require Import UniMath.CategoryTheory.DisplayedCats.Auxiliary. |
| 39 | +Require Import UniMath.CategoryTheory.DisplayedCats.Core. |
| 40 | +Require Import UniMath.CategoryTheory.DisplayedCats.Constructions. |
| 41 | +Require Import UniMath.CategoryTheory.DisplayedCats.Fibrations. |
| 42 | + |
| 43 | +Require Import UniMath.CategoryTheory.Subcategory.Core. |
| 44 | +Require Import UniMath.CategoryTheory.Subcategory.Full. |
| 45 | + |
| 46 | +Local Notation MONAD := (Monad SET). |
| 47 | +Local Notation MODULE R := (LModule R SET). |
| 48 | + |
| 49 | +(** |
| 50 | +The command: |
| 51 | +
|
| 52 | +Check (x ::= y) |
| 53 | +
|
| 54 | +succeeds if and only if [x] is convertible to [y] |
| 55 | +
|
| 56 | +*) |
| 57 | +Notation "x ::= y" := ((idpath _ : x = y) = idpath _) (at level 70, no associativity). |
| 58 | + |
| 59 | +Fail Check (true ::= false). |
| 60 | +Check (true ::= true). |
| 61 | + |
| 62 | +(** ******************* |
| 63 | +
|
| 64 | + Definition of a signature |
| 65 | +
|
| 66 | +The more detailed definition can be found in Signature/Signature.v |
| 67 | +
|
| 68 | + ****** *) |
| 69 | + |
| 70 | +Check ( |
| 71 | + signature_data (C := SET) ::= |
| 72 | + (** a signature assigns to each monad a module over it *) |
| 73 | + ∑ F : (∏ R : MONAD, MODULE R), |
| 74 | + (** a signature sends monad morphisms on module morphisms *) |
| 75 | + ∏ (R S : MONAD) (f : Monad_Mor R S), LModule_Mor _ (F R) (pb_LModule f (F S)) |
| 76 | + ). |
| 77 | + |
| 78 | +(** Functoriality for signatures: |
| 79 | +- [signature_idax] means that # F id = id |
| 80 | +- [signature_compax] means that # F (f o g) = # F f o # F g |
| 81 | + *) |
| 82 | +Check (∏ (F : signature_data (C:= SET)), |
| 83 | + is_signature F ::= |
| 84 | + (signature_idax F × signature_compax F)). |
| 85 | + |
| 86 | +Check (signature SET ::= ∑ F : signature_data, is_signature F). |
| 87 | + |
| 88 | + |
| 89 | +Local Notation SIGNATURE := (signature SET). |
| 90 | +(** ******************* |
| 91 | +
|
| 92 | + Definition of a model of a signature |
| 93 | +
|
| 94 | +The more detailed definitions of models can be found in Signatures/Signature.v |
| 95 | +The more detailed definitions of model morphisms can be found in SoftSignatures/ModelCat.v |
| 96 | +
|
| 97 | + ****** *) |
| 98 | + |
| 99 | +(** The tautological module R over the monad R *) |
| 100 | +Local Notation Θ := tautological_LModule. |
| 101 | + |
| 102 | +Check (∏ (R : MONAD), (Θ R : functor _ _) ::= R). |
| 103 | + |
| 104 | +(** |
| 105 | +A model of a signature S is a monad R with a module morphism from S R to R, called an action. |
| 106 | +*) |
| 107 | +Check (∏ (S : SIGNATURE), model S ::= |
| 108 | + ∑ R : MONAD, LModule_Mor R (S R) (Θ R)). |
| 109 | + |
| 110 | +(** the action of a model *) |
| 111 | +Check (∏ (S : SIGNATURE) (M : model S), |
| 112 | + model_τ M ::= pr2 M). |
| 113 | + |
| 114 | +(** |
| 115 | +Given a signature S, a model morphism is a monad morphism commuting with the action |
| 116 | +*) |
| 117 | +Check (∏ (S : SIGNATURE) (M N : model S), |
| 118 | + rep_fiber_mor M N ::= |
| 119 | + ∑ g:Monad_Mor M N, |
| 120 | + ∏ c : SET, model_τ M c · g c = ((#S g)%ar:nat_trans _ _) c · model_τ N c ). |
| 121 | + |
| 122 | + |
| 123 | +Local Notation "R →→ S" := (rep_fiber_mor R S) (at level 6). |
| 124 | + |
| 125 | + |
| 126 | +(** The category of 1-models *) |
| 127 | +Check (∏ (S : SIGNATURE), |
| 128 | + rep_fiber_precategory S ::= |
| 129 | + (precategory_data_pair |
| 130 | + (** the object and morphisms of the category *) |
| 131 | + (precategory_ob_mor_pair (model S) rep_fiber_mor) |
| 132 | + (** the identity morphism *) |
| 133 | + (λ R, rep_fiber_id R) |
| 134 | + (** Composition *) |
| 135 | + (λ M N O , rep_fiber_comp) |
| 136 | + ,, |
| 137 | + (** This second component is a proof that the axioms of categories are satisfied *) |
| 138 | + is_precategory_rep_fiber_precategory_data S) |
| 139 | + ). |
| 140 | + |
| 141 | + |
| 142 | +(** ******************* |
| 143 | +
|
| 144 | + Definition of an S-signature, or that of a signature over a 1-signature S |
| 145 | +
|
| 146 | +The more detailed definitions can be found in SoftSignatures/SignatureOver.v |
| 147 | +
|
| 148 | + ****** *) |
| 149 | + |
| 150 | +(** |
| 151 | +a signature over a 1-sig S assigns functorially to any model of S a module over the underlying monad. |
| 152 | +*) |
| 153 | +Check (∏ (S : SIGNATURE), |
| 154 | + signature_over S ::= |
| 155 | + ∑ F : |
| 156 | + (** model -> module over the monad *) |
| 157 | + (∑ F : (∏ R : model S, MODULE R), |
| 158 | + (** model morphism -> module morphism *) |
| 159 | + ∏ (R S : model S) (f : R →→ S), |
| 160 | + LModule_Mor _ (F R) (pb_LModule f (F S))), |
| 161 | + |
| 162 | + (** functoriality conditions (see SignatureOver.v) *) |
| 163 | + is_signature_over S F). |
| 164 | + |
| 165 | +Local Notation "F ⟹ G" := (signature_over_Mor _ F G) (at level 39). |
| 166 | + |
| 167 | +(** |
| 168 | +a morphism of oversignature is a natural transformation |
| 169 | +*) |
| 170 | +Check (∏ (S : SIGNATURE) |
| 171 | + (F F' : signature_over S), |
| 172 | + signature_over_Mor S F F' ::= |
| 173 | + (** a family of module morphism from F R to F' R for any model R *) |
| 174 | + ∑ (f : (∏ R : model S, LModule_Mor R (F R) (F' R))), |
| 175 | + (** subject to naturality conditions (see SignatureOver.v for the full definition) *) |
| 176 | + is_signature_over_Mor S F F' f |
| 177 | + ). |
| 178 | + |
| 179 | +(** Definition of an oversignature which preserve epimorphisms in the category of natural transformations |
| 180 | +(Cf SoftEquations/quotientequation.v *) |
| 181 | +Check (∏ (S : SIGNATURE) |
| 182 | + (F : signature_over S), |
| 183 | + isEpi_overSig F ::= |
| 184 | + ∏ R S (f : R →→ S), |
| 185 | + isEpi (C := [SET, SET]) (f : nat_trans _ _) -> |
| 186 | + isEpi (C := [SET, SET]) (# F f : nat_trans _ _)%sigo |
| 187 | + ). |
| 188 | + |
| 189 | +(** Definition of a soft-over signature (SoftEquations/quotientequation.v) |
| 190 | +
|
| 191 | +It is a signature Σ such that for any model R, and any family of model morphisms |
| 192 | +(f_j : R --> d_j), the following diagram can be completed in the category |
| 193 | +of natural transformations: |
| 194 | +
|
| 195 | +<<< |
| 196 | + Σ(f_j) |
| 197 | + Σ(R) -----------> Σ(d_j) |
| 198 | + | |
| 199 | + | |
| 200 | + | |
| 201 | + Σ(π)| |
| 202 | + | |
| 203 | + V |
| 204 | + Σ(S) |
| 205 | +
|
| 206 | +>>> |
| 207 | +
|
| 208 | +where π : R -> S is the canonical projection (S is R quotiented by the family (f_j)_j |
| 209 | +
|
| 210 | + *) |
| 211 | +Check (∏ (S : SIGNATURE) |
| 212 | + (F : signature_over S) |
| 213 | + (** The axiom of choice is necessary to make quotient monad *) |
| 214 | + (ax_choice : AxiomOfChoice.AxiomOfChoice_surj) |
| 215 | + (** S preserves epimorphisms of monads *) |
| 216 | + (isEpi_sig : ∏ (R R' : MONAD) |
| 217 | + (f : Monad_Mor R R'), |
| 218 | + (isEpi (C:= [SET,SET]) (f : nat_trans _ _) -> |
| 219 | + isEpi (C:= [SET,SET]) ((#S f)%ar : nat_trans _ _))), |
| 220 | + isSoft ax_choice isEpi_sig F |
| 221 | + ::= |
| 222 | + (∏ (R : model S) |
| 223 | + (J : UU)(d : J -> (model S))(f : ∏ j, R →→ (d j)) |
| 224 | + X (x y : (F R X : hSet)) |
| 225 | + (pi := projR_rep S isEpi_sig ax_choice d f), |
| 226 | + (∏ j, (# F (f j))%sigo X x = (# F (f j))%sigo X y ) |
| 227 | + -> (# F pi X x)%sigo = |
| 228 | + (# F pi X y)%sigo ) |
| 229 | + ). |
| 230 | + |
| 231 | + |
| 232 | + |
| 233 | + |
| 234 | + |
| 235 | +(* ********** |
| 236 | +
|
| 237 | +Definition of Equations. See SoftEquations/Equation.v |
| 238 | +
|
| 239 | +************ **) |
| 240 | + |
| 241 | +(** An equation over a signature S is a pair of two signatures over S, and a signature over morphism between them *) |
| 242 | +Check (∏ (S : SIGNATURE), |
| 243 | + equation (Sig := S) ::= |
| 244 | + ∑ (S1 S2 : signature_over S), S1 ⟹ S2 × S1 ⟹ S2). |
| 245 | + |
| 246 | +(** Soft equation: the domain must be an epi over-signature, and the target |
| 247 | + must be soft (SoftEquations/quotientequation) |
| 248 | + *) |
| 249 | +Check (∏ (S : SIGNATURE) |
| 250 | + (** The axiom of choice is necessary to make quotient monad *) |
| 251 | + (ax_choice : AxiomOfChoice.AxiomOfChoice_surj) |
| 252 | + (** S preserves epimorphisms of monads *) |
| 253 | + (isEpi_sig : ∏ (R R' : MONAD) |
| 254 | + (f : Monad_Mor R R'), |
| 255 | + (isEpi (C:= [SET,SET]) (f : nat_trans _ _) -> |
| 256 | + isEpi (C:= [SET,SET]) ((#S f)%ar : nat_trans _ _))), |
| 257 | + |
| 258 | + soft_equation ax_choice isEpi_sig ::= |
| 259 | + ∑ (e : equation), isSoft ax_choice isEpi_sig (pr1 (pr2 e)) × isEpi_overSig (pr1 e)). |
| 260 | +(** |
| 261 | +Definition of the category of 2-models of a 1-signature with a family of equation. |
| 262 | +
|
| 263 | +It is the full subcategory of 1-models satisfying all the equations |
| 264 | +
|
| 265 | +See SoftEquations/Equation.v for details |
| 266 | +
|
| 267 | +*) |
| 268 | + |
| 269 | +(** a model R satifies the equations if the R-component of the two over-signature morphisms are equal for any |
| 270 | + equation of the family *) |
| 271 | +Check (∏ (S : SIGNATURE) |
| 272 | + (** a family of equations indexed by O *) |
| 273 | + O (e : O -> equation (Sig := S)) |
| 274 | + (R : model S) |
| 275 | + , |
| 276 | + satisfies_all_equations_hp e R ::= |
| 277 | + ( |
| 278 | + (∏ (o : O), |
| 279 | + (** the first half-equation of the equation [e o] *) |
| 280 | + pr1 (pr2 (pr2 (e o))) R = |
| 281 | + (** the second half-equation of the equation [e o] *) |
| 282 | + pr2 (pr2 (pr2 (e o))) R) |
| 283 | + ,, |
| 284 | + (** this second component is a proof that this predicate is hProp *) |
| 285 | + _) |
| 286 | + ). |
| 287 | + |
| 288 | +(** The category of 2-models is the full subcategory of the category [rep_fiber_category S] |
| 289 | + satisfying all the equations *) |
| 290 | +Check (∏ (S : SIGNATURE) |
| 291 | + (** a family of equations indexed by O *) |
| 292 | + O (e : O -> equation (Sig := S)), |
| 293 | + |
| 294 | + precategory_model_equations e ::= |
| 295 | + full_sub_precategory (C := rep_fiber_precategory S) |
| 296 | + (satisfies_all_equations_hp e)). |
| 297 | + |
| 298 | + |
| 299 | +(** *********************** |
| 300 | +
|
| 301 | +Our main result : if a 1-signature Σ generates a syntax, then the 2-signature over Σ |
| 302 | +consisting of any family of soft equations over Σ also generates a syntax |
| 303 | +(SoftEquations/InitialEquationRep.v) |
| 304 | +
|
| 305 | +*) |
| 306 | +Check (soft_equations_preserve_initiality : |
| 307 | + (** The axiom of choice is needed for quotient monads/models *) |
| 308 | + ∏ (choice : AxiomOfChoice.AxiomOfChoice_surj) |
| 309 | + (** The 1-signature *) |
| 310 | + (Sig : SIGNATURE) |
| 311 | + (** The 1-signature must be an epi-signature *) |
| 312 | + (epiSig : ∏ (R S : Monad SET) (f : Monad_Mor R S), |
| 313 | + isEpi (C := [SET, SET]) (f : nat_trans R S) |
| 314 | + → isEpi (C := [SET, SET]) |
| 315 | + ((# Sig)%ar f : nat_trans (Sig R) (pb_LModule f (Sig S)))) |
| 316 | + (** A family of equations *) |
| 317 | + (O : UU) (eq : O → soft_equation choice epiSig), |
| 318 | + (** If the category of 1-models has an initial object, .. *) |
| 319 | + Initial (rep_fiber_category Sig) |
| 320 | + (** .. then the category of 2-models has an initial object *) |
| 321 | + → Initial (precategory_model_equations (λ x : O, eq x))). |
| 322 | + |
| 323 | + |
0 commit comments