From 434dfe28b5f297c41c12d2272ddadf250d810f40 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Na=C3=AFm=20Favier?= Date: Thu, 1 Aug 2024 04:55:24 +0200 Subject: [PATCH 1/9] chore: misc nonsense (#419) --- src/1Lab/Path.lagda.md | 2 +- src/Cat/Displayed/Total.lagda.md | 3 +++ src/Data/Nat/Base.lagda.md | 2 +- support/sort-imports.hs | 37 ++++++++++++++++---------------- 4 files changed, 23 insertions(+), 21 deletions(-) diff --git a/src/1Lab/Path.lagda.md b/src/1Lab/Path.lagda.md index 474002dce..3ccdd94ff 100644 --- a/src/1Lab/Path.lagda.md +++ b/src/1Lab/Path.lagda.md @@ -771,7 +771,7 @@ In Cubical Agda, the relevant *primitive* is the function for now. To start with, this is where paths show their difference from the notion of equality in set-level type theories: it says that we have a function from paths $p : A \is B$ to functions $A \to B$. However, -it's *not* the case that every $p, q : A \to B$ gives back the *same* +it's *not* the case that every $p, q : A \is B$ gives back the *same* function $A \to B$. Which function you get depends on (and determines) the path you put in! diff --git a/src/Cat/Displayed/Total.lagda.md b/src/Cat/Displayed/Total.lagda.md index 0292e2160..e18edc748 100644 --- a/src/Cat/Displayed/Total.lagda.md +++ b/src/Cat/Displayed/Total.lagda.md @@ -277,5 +277,8 @@ module _ {o ℓ o' ℓ'} {B : Precategory o ℓ} {E : Displayed B o' ℓ'} where → {X Y : Total E} ⦃ i : Funlike (Hom (X .fst) (Y .fst)) A B ⦄ → Funlike (Total-hom E X Y) A B Funlike-Total-hom ⦃ i ⦄ .Funlike._#_ f x = f .Total-hom.hom # x + + H-Level-Total-hom' : ∀ {X Y} {n} → H-Level (Total-hom E X Y) (2 + n) + H-Level-Total-hom' = H-Level-Total-hom E ``` --> diff --git a/src/Data/Nat/Base.lagda.md b/src/Data/Nat/Base.lagda.md index 9da4870d7..092851e38 100644 --- a/src/Data/Nat/Base.lagda.md +++ b/src/Data/Nat/Base.lagda.md @@ -33,7 +33,7 @@ Nat-elim : ∀ {ℓ} (P : Nat → Type ℓ) → ({n : Nat} → P n → P (suc n)) → (n : Nat) → P n Nat-elim P pz ps zero = pz -Nat-elim P pz ps (suc n) = Nat-elim (λ z → P (suc z)) (ps pz) ps n +Nat-elim P pz ps (suc n) = ps (Nat-elim P pz ps n) iter : ∀ {ℓ} {A : Type ℓ} → Nat → (A → A) → A → A iter zero f = id diff --git a/support/sort-imports.hs b/support/sort-imports.hs index bd686f9aa..0fda7afcd 100755 --- a/support/sort-imports.hs +++ b/support/sort-imports.hs @@ -4,11 +4,12 @@ --package deepseq --package shake -} -{-# LANGUAGE OverloadedStrings #-} +{-# LANGUAGE OverloadedStrings, BlockArguments #-} module Main where import Control.Exception import Control.DeepSeq +import Control.Monad import Data.List (isSuffixOf, sortOn, groupBy, partition) import qualified Data.Text.IO as Text @@ -39,38 +40,36 @@ sortImports path sortImportsCode :: FilePath -> IO () sortImportsCode path = do - putStrLn $ "Sorting Agda file " ++ path lines <- Text.lines <$> Text.readFile path evaluate (rnf lines) + let sorted = sortImpl lines - withFile path WriteMode $ \handle -> do - traverse_ (Text.hPutStrLn handle) (sortImpl lines) + when (lines /= sorted) do + putStrLn $ "Sorting Agda file " ++ path + withFile path WriteMode $ \handle -> do + traverse_ (Text.hPutStrLn handle) (sortImpl lines) sortImportsLiterate :: FilePath -> IO () sortImportsLiterate path = do - putStrLn $ "Sorting Literate Agda file " ++ path lines <- Text.lines <$> Text.readFile path evaluate (rnf lines) + let + (prefix, first_code_rest) = + break ((||) <$> (== "```agda") <*> (== "```")) lines - withFile path WriteMode $ \handle -> do - let - (prefix, first_code_rest) = - break ((||) <$> (== "```agda") <*> (== "```")) lines - - traverse_ (Text.hPutStrLn handle) prefix - - case first_code_rest of - (pre:lines) -> do - Text.hPutStrLn handle pre - + sorted = prefix ++ case first_code_rest of + pre:lines -> let (code, rest) = break ((||) <$> (== "```agda") <*> (== "```")) lines code' = sortImpl code + in pre : code' ++ rest + _ -> first_code_rest - traverse_ (Text.hPutStrLn handle) code' - traverse_ (Text.hPutStrLn handle) rest - _ -> traverse_ (Text.hPutStrLn handle) first_code_rest + when (lines /= sorted) do + putStrLn $ "Sorting Literate Agda file " ++ path + withFile path WriteMode $ \handle -> do + traverse_ (Text.hPutStrLn handle) sorted sortImpl :: [Text.Text] -> [Text.Text] sortImpl lines = sorted ++ emptyLineBefore' mod where From b077f1c4ccea3af16e1fbf9bfa0de2d7cad044b5 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Wed, 31 Jul 2024 21:45:36 -0700 Subject: [PATCH 2/9] Set-projective types and projective objects (#421) # Description This PR defines projective objects, and proves a bunch of properties about them. It also defines set-projective types, and shows that projectivity of propositions + a set-projective dedekind type implies countable choice. This closes #420. ## Checklist Before submitting a merge request, please check the items below: - [x] I've read [the contributing guidelines](https://github.com/plt-amy/1lab/blob/main/CONTRIBUTING.md). - [x] The imports of new modules have been sorted with `support/sort-imports.hs` (or `nix run --experimental-features nix-command -f . sort-imports`). - [x] All new code blocks have "agda" as their language. If your change affects many files without adding substantial content, and you don't want your name to appear on those pages (for example, treewide refactorings or reformattings), start the commit message and PR title with `chore:`. --- src/1Lab/HIT/Truncation.lagda.md | 30 ++- src/Borceux.lagda.md | 32 +++ src/Cat/Diagram/Coproduct/Indexed.lagda.md | 95 +++++++ src/Cat/Diagram/Projective.lagda.md | 289 +++++++++++++++++++++ src/Cat/Diagram/Projective/Strong.lagda.md | 250 ++++++++++++++++++ src/Cat/Diagram/Separator.lagda.md | 114 ++++---- src/Cat/Functor/Morphism.lagda.md | 39 ++- src/Cat/Morphism.lagda.md | 10 +- src/Data/Set/Projective.lagda.md | 252 ++++++++++++++++++ src/Data/Set/Surjection.lagda.md | 24 +- src/bibliography.bibtex | 7 + 11 files changed, 1069 insertions(+), 73 deletions(-) create mode 100644 src/Cat/Diagram/Projective.lagda.md create mode 100644 src/Cat/Diagram/Projective/Strong.lagda.md create mode 100644 src/Data/Set/Projective.lagda.md diff --git a/src/1Lab/HIT/Truncation.lagda.md b/src/1Lab/HIT/Truncation.lagda.md index f3d13da10..c52af5b38 100644 --- a/src/1Lab/HIT/Truncation.lagda.md +++ b/src/1Lab/HIT/Truncation.lagda.md @@ -6,6 +6,7 @@ definition: | + + diff --git a/src/Borceux.lagda.md b/src/Borceux.lagda.md index 43495e248..db73fa866 100644 --- a/src/Borceux.lagda.md +++ b/src/Borceux.lagda.md @@ -21,6 +21,7 @@ open import Cat.Functor.Adjoint.Continuous open import Cat.Functor.Adjoint.Reflective open import Cat.Diagram.Colimit.Universal open import Cat.Diagram.Coproduct.Indexed +open import Cat.Diagram.Projective.Strong open import Cat.Diagram.Separator.Regular open import Cat.Functor.Hom.Representable open import Cat.Instances.Sets.Cocomplete @@ -844,6 +845,37 @@ _ = zero+separating-family→separator ### 4.6 Projectives +::: warning +Borceux uses the term "projective" to refer to [[strong projectives]]. +::: + + + +* Definition 4.6.1: `is-strong-projective`{.Agda} +* Proposition 4.6.2: + Note that there is a slight typo in Borceux here: $\cC(P,-)$ + must preserve [[strong epimorphisms]]. + (⇒) `preserves-strong-epis→strong-projective`{.Agda} + (⇐) `strong-projective→preserves-strong-epis`{.Agda} +* Proposition 4.6.3: `indexed-coproduct-strong-projective`{.Agda} +* Proposition 4.6.4: `retract→strong-projective`{.Agda} +* Definition 4.6.5: `Strong-projectives`{.Agda} +* Proposition 4.6.6: `strong-projective-separating-faily→strong-projectives`{.Agda} +* Proposition 4.6.7: + * (⇒) `zero+indexed-coproduct-strong-projective→strong-projective`{.Agda} + * (⇐) `indexed-coproduct-strong-projective`{.Agda} + ### 4.7 Injective cogenerators ### 4.8 Exercises diff --git a/src/Cat/Diagram/Coproduct/Indexed.lagda.md b/src/Cat/Diagram/Coproduct/Indexed.lagda.md index a1ab98ff7..5e0b4cbee 100644 --- a/src/Cat/Diagram/Coproduct/Indexed.lagda.md +++ b/src/Cat/Diagram/Coproduct/Indexed.lagda.md @@ -3,8 +3,11 @@ open import Cat.Instances.Shape.Terminal open import Cat.Diagram.Pullback open import Cat.Diagram.Initial +open import Cat.Diagram.Zero open import Cat.Univalent open import Cat.Prelude + +open import Data.Dec ``` --> @@ -304,3 +307,95 @@ is-initial→is-disjoint-coproduct {F = F} {i = i} init = is-disjoint where is-disjoint .summands-intersect i j = absurd i is-disjoint .different-images-are-disjoint i j p = absurd i ``` + +## Coproducts and zero objects + +Let $\cC$ be a category with a [[zero object]], and let $\coprod_{i : I} P_i$ +be a coproduct. If $I$ is a [[discrete]] type, then every coproduct +inclusion $\iota_{i} : \cC(P_i, \coprod_{i : I} P_i)$ has a [[retract]]. + + + +First, a useful lemma. Suppose that we have a coproduct $\coprod_{i : I} P_i$ +indexed by a discrete type, and a map $t_i : \cC(P_i, X)$ for some $i : I$. +If there exists maps $f_j : \cC(P_j, X)$ for every $j \neq i$, then we can +obtain a map $\cC(\coprod_{i : I} P_i, X)$. + + +```agda + detect + : ∀ {X} ⦃ Idx-Discrete : Discrete Idx ⦄ + → (i : Idx) → C.Hom (P i) X + → (∀ (j : Idx) → ¬ i ≡ j → C.Hom (P j) X) + → C.Hom ∐P X +``` + +The key idea here is to check if $i = j$ when invoking the universal +property of $\coprod_{i : I} P_i$; if $i = j$ we use $t_i$, +otherwise we use $f_j$. + +```agda + detect {X = X} ⦃ Idx-Discrete ⦄ i tᵢ fⱼ = match probe + module detect where + probe : ∀ (j : Idx) → C.Hom (P j) X + probe j with i ≡? j + ... | yes i=j = subst _ i=j tᵢ + ... | no ¬i=j = fⱼ j ¬i=j + + probe-yes : probe i ≡ tᵢ + probe-yes with i ≡? i + ... | yes i=i = + is-set→subst-refl + (λ j → C.Hom (P j) X) + (Discrete→is-set Idx-Discrete) + i=i tᵢ + ... | no ¬i=i = absurd (¬i=i refl) + + probe-no : ∀ j → (¬i=j : ¬ (i ≡ j)) → probe j ≡ fⱼ j ¬i=j + probe-no j ¬i=j with i ≡? j + ... | yes i=j = absurd (¬i=j i=j) + ... | no _ = ap (fⱼ j) prop! +``` + +Moreover, we observe that our newly created map interacts nicely +with the inclusions into the coproduct. + +```agda + detect-yes + : ∀ {X} ⦃ Idx-Discrete : Discrete Idx ⦄ + → {i : Idx} → {tᵢ : C.Hom (P i) X} + → {fⱼ : ∀ (j : Idx) → ¬ i ≡ j → C.Hom (P j) X} + → detect i tᵢ fⱼ C.∘ ι i ≡ tᵢ + detect-yes = commute ∙ detect.probe-yes _ _ _ + + detect-no + : ∀ {X} ⦃ Idx-Discrete : Discrete Idx ⦄ + → {i : Idx} → {tᵢ : C.Hom (P i) X} + → {fⱼ : ∀ (j : Idx) → ¬ i ≡ j → C.Hom (P j) X} + → ∀ j → (¬i=j : ¬ i ≡ j) → detect i tᵢ fⱼ C.∘ ι j ≡ fⱼ j ¬i=j + detect-no j ¬i=j = commute ∙ detect.probe-no _ _ _ j ¬i=j +``` + +Refocusing our attention back to our original claim, suppose that +$\cC$ has a zero object. This means that there is a canonical choice +of morphism between any two objects, so we can apply our previous +lemma to obtain a retract $\cC(\coprod_{I} P_i, P_i)$. + +```agda + zero→ι-has-retract + : ∀ ⦃ Idx-Discrete : Discrete Idx ⦄ + → Zero C + → ∀ i → C.has-retract (ι i) + zero→ι-has-retract z i = + C.make-retract (detect i C.id (λ _ _ → zero→)) detect-yes + where open Zero z +``` diff --git a/src/Cat/Diagram/Projective.lagda.md b/src/Cat/Diagram/Projective.lagda.md new file mode 100644 index 000000000..12e220d8e --- /dev/null +++ b/src/Cat/Diagram/Projective.lagda.md @@ -0,0 +1,289 @@ +--- +description: | + Projective objects. +--- + +```agda +module Cat.Diagram.Projective + {o ℓ} + (C : Precategory o ℓ) + where +``` + + + +# Projective objects + +:::{.definition #projective-object alias="projective"} +Let $\cC$ be a precategory. An object $P : \cC$ is **projective** +if for every morphism $p : P \to Y$ and [[epimorphism]] $e : X \epi Y$, +there merely exists a $s : P \to X$ such that $e \circ s = p$, as in the +following diagram: + +~~~{.quiver} +\begin{tikzcd} + && X \\ + \\ + P && Y + \arrow["e", two heads, from=1-3, to=3-3] + \arrow["\exists s", dashed, from=3-1, to=1-3] + \arrow["p"', from=3-1, to=3-3] +\end{tikzcd} +~~~ + +More concisely, $P$ is projective if it has the left-lifting property +relative to epimorphisms in $\cC$. +::: + +```agda +is-projective : (P : Ob) → Type _ +is-projective P = + ∀ {X Y} (p : Hom P Y) (e : X ↠ Y) + → ∃[ s ∈ Hom P X ] (e .mor ∘ s ≡ p) +``` + +If we take the perspective of generalized elements, then a projective +object $P$ lets us pick a $P$-element of $X$ from the preimage $e^{-1}(y)$ +of a $P$-element $y : Y$ along every $e : X \epi Y$. This endows $\cC$ with +an internal $P$-relative version of the [[axiom of choice]]. + +This intuition can be made more precise by noticing that every +object of $\cC$ is projective if and only if every epimorphism (merely) +splits. For the forward direction, let $e : X \epi Y$ have a section +$s : Y \to X$, and note that $s \circ p$ factorizes $p$ through $e$. + +```agda +epis-split→all-projective + : (∀ {X Y} → (e : X ↠ Y) → ∥ has-section (e .mor) ∥) + → ∀ {P} → is-projective P +epis-split→all-projective epi-split p e = do + s ← epi-split e + pure (s .section ∘ p , cancell (s .is-section)) + where open has-section +``` + +Conversely, given an epi $X \epi Y$, we can factorize $\id : Y \to Y$ +to get our desired section $s : Y \to X$. + +```agda +all-projective→epis-split + : (∀ {P} → is-projective P) + → ∀ {X Y} → (e : X ↠ Y) → ∥ has-section (e .mor) ∥ +all-projective→epis-split pro e = do + (s , p) ← pro id e + pure (make-section s p) +``` + +This suggests that projective objects are quite hard to come by +in constructive foundations! For the most part, this is true: +constructively, the category of sets has very few projectives[^1], and +many categories inherit their epimorphisms from $\Sets$. However, +it is still possible to have projectives if epis in $\cC$ are extremely +structured, or there are very few maps out of some $P$. + +[^1]: In fact, it is consistent that [the only projective sets are the finite sets]! + +[the only projective sets are the finite sets]: Data.Set.Projective.html#taboos + + + +For instance, observe that every epi splits in a [[pregroupoid]], +so every every object in a pregroupoid must be projective. + +```agda +pregroupoid→all-projective + : is-pregroupoid C + → ∀ {P} → is-projective P +pregroupoid→all-projective pregroupoid = + epis-split→all-projective λ e → + pure (invertible→to-has-section (pregroupoid (e .mor))) +``` + +Likewise, if $\cC$ has an [[initial object]] $\bot : \cC$, then +$\bot$ is projective, as there is a unique map out of $\bot$. + +```agda +module _ (initial : Initial C) where + open Initial initial + + initial-projective : is-projective bot + initial-projective p e = pure (¡ , ¡-unique₂ (e .mor ∘ ¡) p) +``` + +## A functorial definition + +Some authors prefer to define projective objects via a functorial +approach. In particular, an object $P : \cC$ is projective if and only +if the $\hom$-functor $\cC(P,-)$ preserves epimorphisms. + +For the forward direction, recall that in $\Sets$, [[epis are surjective]]. +This means that if $e : X \epi Y$ is an epi in $\cC$, then +$e \circ - : \cC(P,X) \to \cC(P,Y)$ is surjective, as $\cC(P,-)$ preserves +epis. This directly gives us the factorization we need! + +```agda +preserves-epis→projective + : ∀ {P} + → preserves-epis (Hom-from C P) + → is-projective P +preserves-epis→projective {P = P} hom-epi {X = X} {Y = Y} p e = + epi→surjective (el! (Hom P X)) (el! (Hom P Y)) + (e .mor ∘_) + (λ {c} → hom-epi (e .epic) {c = c}) + p +``` + +For the reverse direciton, let $P$ be projective, $f : X \epi Y$ be an epi, +and $g, h : \cC(P, X) \to A$ be a pair of functions into an arbitrary +set $A$ such that $g(f \circ s) = h(f \circ s)$ for any $s : \cC(P, X)$. +To show that $\cC(P,-)$ preserves epis, we must show that $g = h$, which +follows directly from the existence of a lift for every $\cC(P,X)$. + +```agda +projective→preserves-epis + : ∀ {P} + → is-projective P + → preserves-epis (Hom-from C P) +projective→preserves-epis pro {f = f} f-epi g h p = + ext λ k → + rec! + (λ s s-section → + g k ≡˘⟨ ap g s-section ⟩ + g (f ∘ s) ≡⟨ p $ₚ s ⟩ + h (f ∘ s) ≡⟨ ap h s-section ⟩ + h k ∎) + (pro k (record { epic = f-epi })) +``` + +## Closure of projectives + +Projective objects are equipped with a mapping-out property, so they +tend to interact nicely with other constructions that also have a +mapping-out property. For instance, f $P$ and $Q$ are both projective, +then their [[coproduct]] $P + Q$ is projective (if it exists). + +```agda +coproduct-projective + : ∀ {P Q P+Q} {ι₁ : Hom P P+Q} {ι₂ : Hom Q P+Q} + → is-projective P + → is-projective Q + → is-coproduct C ι₁ ι₂ + → is-projective P+Q +coproduct-projective {ι₁ = ι₁} {ι₂ = ι₂} P-pro Q-pro coprod p e = do + (s₁ , s₁-factor) ← P-pro (p ∘ ι₁) e + (s₂ , s₂-factor) ← Q-pro (p ∘ ι₂) e + pure $ + [ s₁ , s₂ ] , + unique₂ + (pullr []∘ι₁ ∙ s₁-factor) (pullr []∘ι₂ ∙ s₂-factor) + refl refl + where open is-coproduct coprod +``` + +We can extend this result to [[indexed coproducts]], provided that +the indexing type is [[set projective]]. + +```agda +indexed-coproduct-projective + : ∀ {κ} {Idx : Type κ} + → {P : Idx → Ob} {∐P : Ob} {ι : ∀ i → Hom (P i) ∐P} + → is-set-projective Idx ℓ + → (∀ i → is-projective (P i)) + → is-indexed-coproduct C P ι + → is-projective ∐P +indexed-coproduct-projective {P = P} {ι = ι} Idx-pro P-pro coprod {X = X} {Y = Y} p e = do + s ← Idx-pro + (λ i → Σ[ sᵢ ∈ Hom (P i) X ] (e .mor ∘ sᵢ ≡ p ∘ ι i)) (λ i → hlevel 2) + (λ i → P-pro i (p ∘ ι i) e) + pure (match (λ i → s i .fst) , unique₂ (λ i → pullr commute ∙ s i .snd)) + where open is-indexed-coproduct coprod +``` + +Note that this projectivity requirement is required: if projective objects +were closed under arbitrary coproducts, then we would immediately be able +to prove the [[axiom of choice]]: the singleton set is both a projective +object and a [[dense separator]] in $\Sets$, so closure under arbitrary +coproducts would mean that every set is projective, which is precisely +the axiom of choice. + +Putting coproducts aside for a moment, note that projectives are closed +under retracts. This follows by a straightforward bit of algebra. + +```agda +retract→projective + : ∀ {R P} + → is-projective P + → (s : Hom R P) + → has-retract s + → is-projective R +retract→projective P-pro s r p e = do + (t , t-factor) ← P-pro (p ∘ r .retract) e + pure (t ∘ s , pulll t-factor ∙ cancelr (r .is-retract)) +``` + +A nice consequence of this is that if $\cC$ has a [[zero object]] and +a projective coproduct $\coprod_{I} P_i$ indexed by a [[discrete]] type, +then each component of the coproduct is also projective. + +```agda +zero+indexed-coproduct-projective→projective + : ∀ {κ} {Idx : Type κ} ⦃ Idx-Discrete : Discrete Idx ⦄ + → {P : Idx → Ob} {∐P : Ob} {ι : ∀ i → Hom (P i) ∐P} + → Zero C + → is-indexed-coproduct C P ι + → is-projective ∐P + → ∀ i → is-projective (P i) +``` + +This follows immediately from the fact that if $\cC$ has a zero object +and $\coprod_{I} P_i$ is indexed by a discrete type, then each coproduct +inclusion has a retract. + +```agda +zero+indexed-coproduct-projective→projective {ι = ι} z coprod ∐P-pro i = + retract→projective ∐P-pro (ι i) $ + zero→ι-has-retract C coprod z i +``` + +## Enough projectives + +A category $\cC$ is said to have **enough projectives** if for +object $X : \cC$ there is some $P \epi X$ with $P$ projective. +We will refer to these projectives as **projective presentations** +of $X$. + +Note that there are two variations on this condition: one where +there *merely* exists a projective presentation for every $X$, and +another where those presentations are provided as structure. We prefer +to work with the latter, as it tends to be less painful to work with. + +```agda +record Projectives : Type (o ⊔ ℓ) where + field + Pro : Ob → Ob + present : ∀ {X} → Pro X ↠ X + projective : ∀ {X} → is-projective (Pro X) +``` diff --git a/src/Cat/Diagram/Projective/Strong.lagda.md b/src/Cat/Diagram/Projective/Strong.lagda.md new file mode 100644 index 000000000..c75be2694 --- /dev/null +++ b/src/Cat/Diagram/Projective/Strong.lagda.md @@ -0,0 +1,250 @@ +--- +description: | + Strong projective objects. +--- + +```agda +module Cat.Diagram.Projective.Strong + {o ℓ} + (C : Precategory o ℓ) + where +``` + + + +# Strong projective objects + +:::{.definition #strong-projective-object alias="strong-projective"} +Let $\cC$ be a precategory. An object $P : \cC$ is a +**strong projective object** if it has the left-lifting property against +[[strong epimorphisms]]. + +More explicitly, $P$ is a strong projective object if for every +morphism $p : P \to Y$ and strong epi $e : X \epi Y$, there merely exists +a $s : P \to X$ such that $e \circ s = p$, as in the following diagram: + +~~~{.quiver} +\begin{tikzcd} + && X \\ + \\ + P && Y + \arrow["e", two heads, from=1-3, to=3-3] + \arrow["\exists s", dashed, from=3-1, to=1-3] + \arrow["p"', from=3-1, to=3-3] +\end{tikzcd} +~~~ +::: + +```agda +is-strong-projective : (P : Ob) → Type _ +is-strong-projective P = + ∀ {X Y} (p : Hom P Y) (e : Hom X Y) + → is-strong-epi e + → ∃[ s ∈ Hom P X ] (e ∘ s ≡ p) +``` + +::: warning +Being a strong projective object is actually a weaker condition than +being a [[projective object]]: strong projectives only need to lift +against strong epis, whereas projectives need to lift against *all* epis. +::: + +```agda +projective→strong-projective + : ∀ {P} + → is-projective P + → is-strong-projective P +projective→strong-projective pro p e e-strong = + pro p (record { mor = e ; epic = e-strong .fst }) +``` + +## A functorial definition + +Like their non-strong counterparts, we can give a functorial definition of +strong projectives. In particular, an object $P : \cC$ is a strong projective +if and only if the $\hom$-functor $\cC(P,-)$ preserves strong epimorphisms. + +```agda +preserves-strong-epis→strong-projective + : ∀ {P} + → preserves-strong-epis (Hom-from C P) + → is-strong-projective P + +strong-projective→preserves-strong-epis + : ∀ {P} + → is-strong-projective P + → preserves-strong-epis (Hom-from C P) +``` + +
+These proofs are essentially the same as the corresponding +ones for projective objects, so we omit the details. + +```agda +preserves-strong-epis→strong-projective {P = P} hom-epi {X = X} {Y = Y} p e e-strong = + epi→surjective (el! (Hom P X)) (el! (Hom P Y)) + (e ∘_) + (λ {c} → hom-epi e-strong .fst {c = c}) + p + +strong-projective→preserves-strong-epis {P = P} pro {X} {Y} {f = f} f-strong = + surjective→strong-epi (el! (Hom P X)) (el! (Hom P Y)) (f ∘_) $ λ p → + pro p f f-strong +``` +
+ +## Closure of strong projectives + +Like projective objects, strong projectives are closed under coproducts +indexed by [[set-projective]] types and retracts. + +```agda +indexed-coproduct-strong-projective + : ∀ {κ} {Idx : Type κ} + → {P : Idx → Ob} {∐P : Ob} {ι : ∀ i → Hom (P i) ∐P} + → is-set-projective Idx ℓ + → (∀ i → is-strong-projective (P i)) + → is-indexed-coproduct C P ι + → is-strong-projective ∐P + +retract→strong-projective + : ∀ {R P} + → is-strong-projective P + → (s : Hom R P) + → has-retract s + → is-strong-projective R +``` + +
+These proofs are more or less identical to the corresponding +ones for projective objects. + +```agda +indexed-coproduct-strong-projective {P = P} {ι = ι} Idx-pro P-pro coprod {X = X} {Y = Y} p e e-strong = do + s ← Idx-pro + (λ i → Σ[ sᵢ ∈ Hom (P i) X ] (e ∘ sᵢ ≡ p ∘ ι i)) (λ i → hlevel 2) + (λ i → P-pro i (p ∘ ι i) e e-strong) + pure (match (λ i → s i .fst) , unique₂ (λ i → pullr commute ∙ s i .snd)) + where open is-indexed-coproduct coprod + +retract→strong-projective P-pro s r p e e-strong = do + (t , t-factor) ← P-pro (p ∘ r .retract) e e-strong + pure (t ∘ s , pulll t-factor ∙ cancelr (r .is-retract)) +``` +
+ +Moreover, if $\cC$ has a [[zero object]] and a strong projective +coproduct $\coprod_{I} P_i$ indexed by a [[discrete]] type, then +each component of the coproduct is a strong projective. + +```agda +zero+indexed-coproduct-strong-projective→strong-projective + : ∀ {κ} {Idx : Type κ} ⦃ Idx-Discrete : Discrete Idx ⦄ + → {P : Idx → Ob} {∐P : Ob} {ι : ∀ i → Hom (P i) ∐P} + → Zero C + → is-indexed-coproduct C P ι + → is-strong-projective ∐P + → ∀ i → is-strong-projective (P i) +``` + +
+Following the general theme, the proof is identical +to the non-strong case. + +```agda +zero+indexed-coproduct-strong-projective→strong-projective {ι = ι} z coprod ∐P-pro i = + retract→strong-projective ∐P-pro (ι i) $ + zero→ι-has-retract C coprod z i +``` +
+ +## Enough strong projectives + +A category $\cC$ is said to have **enough strong projectives** if for +object $X : \cC$ there is some strong epi $P \epi X$ with $P$ strong projective. +We will refer to these projectives as **projective presentations** +of $X$. + +Note that there are two variations on this condition: one where +there *merely* exists a strong projective presentation for every $X$, and +another where those presentations are provided as structure. We prefer +to work with the latter, as it tends to be less painful to work with. + +```agda +record Strong-projectives : Type (o ⊔ ℓ) where + field + Pro : Ob → Ob + present : ∀ {X} → Hom (Pro X) X + present-strong-epi : ∀ {X} → is-strong-epi (present {X}) + projective : ∀ {X} → is-strong-projective (Pro X) +``` + + + +If $\cC$ has set-indexed coproducts, and $P_i$ is a [[strong separating family]] +with each $P_i$ a strong projective, then $\cC$ has enough strong projectives +if $\Sigma(i : Idx) (\cC(P_i, X))$ is a set-projective type. + +```agda + strong-projective-separating-faily→strong-projectives + : ∀ {Idx : Set ℓ} {Pᵢ : ∣ Idx ∣ → Ob} + → (∀ X → is-set-projective (Σ[ i ∈ ∣ Idx ∣ ] (Hom (Pᵢ i) X)) ℓ) + → (∀ i → is-strong-projective (Pᵢ i)) + → is-strong-separating-family Idx Pᵢ + → Strong-projectives +``` + +The hypotheses of this theorem basically give the game away: by definition, +there is a strong epimorphism $\coprod_{\Sigma(i : I) \cC(P_i, X)} S_i \to X$ +for every $X$. Moreover, $\Sigma(i : I) \cC(P_i, X)$ is set-projective, +so the corresponding coproduct is a strong projective. + +```agda + strong-projective-separating-faily→strong-projectives + {Idx} {Pᵢ} Idx-pro Pᵢ-pro strong-sep = strong-projectives where + open Strong-projectives + + strong-projectives : Strong-projectives + strong-projectives .Pro X = + ∐! (Σ[ i ∈ ∣ Idx ∣ ] (Hom (Pᵢ i) X)) (Pᵢ ⊙ fst) + strong-projectives .present {X = X} = + ∐!.match (Σ[ i ∈ ∣ Idx ∣ ] (Hom (Pᵢ i) X)) (Pᵢ ⊙ fst) snd + strong-projectives .present-strong-epi = + strong-sep + strong-projectives .projective {X = X} = + indexed-coproduct-strong-projective + (Idx-pro X) + (Pᵢ-pro ⊙ fst) + (∐!.has-is-ic (Σ[ i ∈ ∣ Idx ∣ ] (Hom (Pᵢ i) X)) (Pᵢ ⊙ fst)) +``` diff --git a/src/Cat/Diagram/Separator.lagda.md b/src/Cat/Diagram/Separator.lagda.md index 0a567b6d6..2679bae93 100644 --- a/src/Cat/Diagram/Separator.lagda.md +++ b/src/Cat/Diagram/Separator.lagda.md @@ -328,89 +328,73 @@ equalisers+jointly-conservative→separating-family ``` + +```agda +``` + Our next result lets us relate separating objects and separating families. Clearly, a separating object yields a separating family; when does the -converse hold? One possible scenario is when: +converse hold? -1. The separating family $S_i$ is indexed by a [[discrete]] type. -2. $\cC$ has $I$-indexed coproducts. -3. Every $\hom$-set $\cC(S_i, X)$ has a distinguished inhabitant. +One possible scenario is that there is an object $X : \cC$ such that +every $S_i$ comes equipped with a section/retraction pair +$f_i : S_i \rightleftarrows X : r_i$. To see why, let $g, h : \cC(A, B)$ +and suppose that $g \circ e = h \circ e$ for every $e : \cC(X, A)$. +As $S_i$ is a separating family, it suffices to show that $g \circ e_i = h \circ e_i$ +for every $i : I$, $e_i : \cC(S_i, A)$. -```agda -module _ - {κ} {Idx : Type κ} {sᵢ : Idx → Ob} - ⦃ Idx-Discrete : Discrete Idx ⦄ - (coprods : has-coproducts-indexed-by C Idx) - where - open Indexed-coproducts-by C coprods +Next, note that $g \circ e_i = g \circ e_i \circ r_i \circ f_i$, as +$f_i$ and $r_i$ are a section/retraction pair. Moreover, +$e_i \circ r_i : \cC(X, A)$, so - hom-inhabited+separating-family→separator - : (∀ i x → Hom (sᵢ i) x) - → is-separating-family sᵢ - → is-separator (ΣF sᵢ) -``` +$$g \circ e_i \circ r_i \circ f_i = h \circ e_i \circ r_i \circ f_i$$ -We shall show that the coproduct $\coprod_{i : I} S_i$ is a separating object. -Suppose that $f, g : \cC(X,Y)$ such that $f \circ e = g \circ e$ for any -$\cC(\coprod_{i : I} S_i, X)$. $S_i$ is a separating family, so we can -attempt to show that $f = g$ by showing that $f \circ e_i = g \circ e_i$ -for every $e_i : \cC(S_i, X)$. - -This is where we need to start using our somewhat contrived assumptions. -The crux of the problem is that we need to somehow turn our proof obligation -into a question of equality involving $\cC(\coprod_{i : I} S_i, X)$. In -particular, if we could factorize $e_i$ into a coproduct injection $\iota$ -followed by a map $u$ out of the coproduct, then our hypothesis would let us -deduce that $f \circ u \circ \iota = g \circ u \circ \iota$. Unfortunatelly, -cooking up a factorization is a bit tricky! +by our hypothesis. Finally, we can use the fact that $f_i$ and $r_i$ +are a section/retraction pair to observe that $g \circ e_i = f \circ e_i$, +completeing the proof ```agda - hom-inhabited+separating-family→separator hom-default separate {x = x} {y = y} {f = f} {g = g} p = - separate λ {i} eᵢ → - f ∘ eᵢ ≡˘⟨ ap (f ∘_) (ι-commute _ ∙ detect-yes i eᵢ) ⟩ - f ∘ match sᵢ (detect i eᵢ) ∘ ι sᵢ i ≡⟨ extendl (p _) ⟩ - g ∘ match sᵢ (detect i eᵢ) ∘ ι sᵢ i ≡⟨ ap (g ∘_) (ι-commute _ ∙ detect-yes i eᵢ) ⟩ - g ∘ eᵢ ∎ - where +retracts+separating-family→separator + : ∀ {κ} {Idx : Type κ} {sᵢ : Idx → Ob} {x : Ob} + → is-separating-family sᵢ + → (fᵢ : ∀ i → Hom (sᵢ i) x) + → (∀ i → has-retract (fᵢ i)) + → is-separator x +retracts+separating-family→separator separate fᵢ r {f = g} {g = h} p = + separate λ {i} eᵢ → + g ∘ eᵢ ≡˘⟨ pullr (cancelr (r i .is-retract)) ⟩ + (g ∘ eᵢ ∘ r i .retract) ∘ fᵢ i ≡⟨ ap₂ _∘_ (p (eᵢ ∘ r i .retract)) refl ⟩ + (h ∘ eᵢ ∘ r i .retract) ∘ fᵢ i ≡⟨ pullr (cancelr (r i .is-retract)) ⟩ + h ∘ eᵢ ∎ ``` -The key idea is that we can extend a single morphism $e_i : \cC(S_i, X)$ to a -family of morphisms $e_{i}^* : (j : I) \to \cC(S_j, X)$ by checking if $i = j$; if it does, then -we simply return $e_i$; if not, then we use the distinguished inhabitant -of $\cC(S_j, X)$. By definition, $e_{i}^*(i) = e_i$, so this construction -provides the required factorization. +We can immediately use this result to note that a separating family +$S_i$ can be turned into a separating object, provided that: + +1. The separating family $S_i$ is indexed by a [[discrete]] type. +2. The coproduct of $\coprod_{I} S_i$ exists. +3. $\cC$ has a zero object. -```agda - detect : ∀ {x} (i : Idx) (eᵢ : Hom (sᵢ i) x) → (j : Idx) → Hom (sᵢ j) x - detect i eᵢ j with i ≡? j - ... | yes i=j = subst _ i=j eᵢ - ... | no _ = hom-default j _ - detect-yes : ∀ {x} (i : Idx) (eᵢ : Hom (sᵢ i) x) → detect i eᵢ i ≡ eᵢ - detect-yes {x = x} i eᵢ with i ≡? i - ... | yes i=i = - is-set→subst-refl - (λ i → Hom (sᵢ i) x) - (Discrete→is-set Idx-Discrete) - i=i eᵢ - ... | no ¬i=i = absurd (¬i=i refl) +```agda +zero+separating-family→separator + : ∀ {κ} {Idx : Type κ} {sᵢ : Idx → Ob} {∐s : Ob} {ι : ∀ i → Hom (sᵢ i) ∐s} + → ⦃ Idx-Discrete : Discrete Idx ⦄ + → Zero C + → is-separating-family sᵢ + → is-indexed-coproduct C sᵢ ι + → is-separator ∐s ``` -In particular, if $\cC$ meets conditions (1) and (2) and has a [[zero object]], -then $\cC$ has a separating object. +This follows immediately from the fact that coproduct inclusions have +retracts when hypotheses (1) and (3) hold. ```agda - zero+separating-family→separator - : Zero C - → is-separating-family sᵢ - → is-separator (ΣF sᵢ) - zero+separating-family→separator z separates = - hom-inhabited+separating-family→separator - (λ _ _ → Zero.zero→ z) - separates +zero+separating-family→separator {ι = ι} z separates coprod = + retracts+separating-family→separator separates ι $ + zero→ι-has-retract C coprod z ``` - # Dense separators As noted in the previous sections, separating objects categorify diff --git a/src/Cat/Functor/Morphism.lagda.md b/src/Cat/Functor/Morphism.lagda.md index 591031e5d..e2b023082 100644 --- a/src/Cat/Functor/Morphism.lagda.md +++ b/src/Cat/Functor/Morphism.lagda.md @@ -6,6 +6,7 @@ description: | + +Likewise, a functor $F : \cC \to \cD$ **reflects** $H$-morphisms +if $F(f) \in H$ implies that $f \in H$. + +```agda +reflects-monos : Type _ +reflects-monos = + ∀ {a b : 𝒞.Ob} {f : 𝒞.Hom a b} → 𝒟.is-monic (F₁ f) → 𝒞.is-monic f + +reflects-epis : Type _ +reflects-epis = + ∀ {a b : 𝒞.Ob} {f : 𝒞.Hom a b} → 𝒟.is-epic (F₁ f) → 𝒞.is-epic f +``` + + ## Faithful functors [[Faithful functors]] reflect [[monomorphisms]] and [[epimorphisms]]. diff --git a/src/Cat/Morphism.lagda.md b/src/Cat/Morphism.lagda.md index 9cde3bd4c..492df4d99 100644 --- a/src/Cat/Morphism.lagda.md +++ b/src/Cat/Morphism.lagda.md @@ -219,7 +219,7 @@ has-section→epic {f = f} f-sect g h p = h ∎ ``` -## Retracts +## Retracts {defines="retract"} A morphism $r : A \to B$ is a retract of another morphism $s : B \to A$ when $r \cdot s = id$. Note that this is the same equation involved @@ -744,6 +744,13 @@ inverses→from-has-retract → Inverses f g → has-retract g inverses→from-has-retract {f = f} inv .retract = f inverses→from-has-retract inv .is-retract = Inverses.invl inv +``` + + Using our lemmas about section/retraction pairs from before, we can show that if $f$ is a monic retract, then $f$ is an diff --git a/src/Data/Set/Projective.lagda.md b/src/Data/Set/Projective.lagda.md new file mode 100644 index 000000000..3f5075fdc --- /dev/null +++ b/src/Data/Set/Projective.lagda.md @@ -0,0 +1,252 @@ +--- +description: | + We define set-projective types, and prove some taboos. +--- + +```agda +module Data.Set.Projective where +``` + +# Set-projective types + + + +:::{.definition #set-projective} +A type $A$ is **set-projective** if we can commute [[propositional truncation]] +past $A$-indexed families. +::: + +```agda +is-set-projective : ∀ {ℓ} (A : Type ℓ) → (κ : Level) → Type _ +is-set-projective A κ = + ∀ (P : A → Type κ) + → (∀ a → is-set (P a)) + → (∀ a → ∥ P a ∥) + → ∥ (∀ a → P a) ∥ +``` + +Intuitively, a type $A$ is set-projective if it validates a sort of +$A$-indexed version of the [[axiom of choice]]. + +If $A$ is a set, then $A$ is set-projective if and only if every +surjection $E \to A$ from a set $E$ splits. + +```agda +surjections-split→set-projective + : ∀ {ℓ ℓ'} {A : Type ℓ} + → is-set A + → (∀ (E : Type (ℓ ⊔ ℓ')) → is-set E + → (f : E → A) → is-surjective f + → ∥ (∀ a → fibre f a) ∥) + → is-set-projective A (ℓ ⊔ ℓ') + +sets-projective→surjections-split + : ∀ {ℓ ℓ'} {A : Type ℓ} + → is-set A + → is-set-projective A (ℓ ⊔ ℓ') + → ∀ {E : Type ℓ'} → is-set E + → (f : E → A) → is-surjective f + → ∥ (∀ a → fibre f a) ∥ +``` + +
+This is essentially a specialization of the of the proof that +the axiom of choice is equivalent to every surjection splitting, so we +will not dwell on the details. + +```agda +surjections-split→set-projective {A = A} A-set surj-split P P-set ∥P∥ = + ∥-∥-map + (Equiv.to (Π-cod≃ (Fibre-equiv P))) + (surj-split (Σ[ x ∈ A ] (P x)) (Σ-is-hlevel 2 A-set P-set) fst λ x → + ∥-∥-map (Equiv.from (Fibre-equiv P x)) (∥P∥ x)) + +sets-projective→surjections-split A-set A-pro E-set f = + A-pro (fibre f) (λ x → fibre-is-hlevel 2 E-set A-set f x) +``` +
+ +## Closure of set-projectivity + +Set-projective types are closed under Σ-types. Suppose that $A : \ty$ is a +set-projective type, and that $B : A \to \ty$ is a family of set-projective +types, and let $P : \Sigma\ A\ B \to \set$ be a family of merely inhabited sets. +Note that $(b : B(a)) \to P(a, b)$ is a $B(a)$-indexed family of merely +inhabited sets for every $a$, so its product must also be inhabited by projectivity +of $B(a)$. Moreover, $A$ is also projective, so $(a : A) \to (b : B(a)) \to P(a, b)$ +is also merely inhabited, as $(b : B(a)) \to P(a, b)$ is an $A$-indexed family of +merely inhabited sets. We can then uncurry this family to finish the proof. + +```agda +Σ-set-projective + : ∀ {A : Type ℓ} {B : A → Type ℓ'} + → is-set-projective A (ℓ' ⊔ ℓ'') + → (∀ a → is-set-projective (B a) ℓ'') + → is-set-projective (Σ[ a ∈ A ] B a) ℓ'' +Σ-set-projective {A = A} {B = B} A-pro B-pro P P-set ∥P∥ = do + ∥-∥-map uncurry $ + A-pro (λ a → ((b : B a) → P (a , b))) (λ a → Π-is-hlevel 2 λ b → P-set (a , b)) λ a → + B-pro a (λ b → P (a , b)) (λ b → P-set (a , b)) λ b → ∥P∥ (a , b) +``` + +Moreover, set-projective types are stable under retracts. Suppose that +we have $f : A \to B, g : B \to A$ with $f \circ g = id$ with $A$ set-projective, +and let $P : B \to \set$ be a family of merely inhabited sets. We can +precompose $P$ with $f$ to obtain an $A$-indexed family of sets whose +product $\Pi (a : A) \to P(f(a))$ must be inhabited via projectivity of $A$. +Moreover, we can precompose again with $g$ to see that $\Pi (b : B) \to P(f(g(b)))$ +is merely inhabited. Finally, $f(g(b)) = b$, so $\Pi (b : B) \to P(b)$ is +merely inhabited. + +```agda +retract→set-projective + : ∀ {A : Type ℓ} {B : Type ℓ'} + → (f : A → B) (g : B → A) + → is-left-inverse f g + → is-set-projective A ℓ'' + → is-set-projective B ℓ'' +retract→set-projective {A = A} {B = B} f g retract A-pro P P-set ∥P∥ = + ∥-∥-map (λ k b → subst P (retract b) (k (g b))) + (A-pro (P ∘ f) (P-set ∘ f) (∥P∥ ∘ f)) +``` + +This gives us a nice proof that set-projectivity is stable under equivalence. + +```agda +Equiv→set-projective + : ∀ {A : Type ℓ} {B : Type ℓ'} + → A ≃ B + → is-set-projective A ℓ'' + → is-set-projective B ℓ'' +Equiv→set-projective f A-pro = + retract→set-projective (Equiv.to f) (Equiv.from f) (Equiv.ε f) A-pro +``` + +By the theorem of [[finite choice]], finite sets are projective. + +```agda +Fin-set-projective : ∀ {n} → is-set-projective (Fin n) ℓ +Fin-set-projective {n = n} P P-set ∥P∥ = finite-choice n ∥P∥ + +finite→set-projective + : {A : Type ℓ} + → Finite A + → is-set-projective A ℓ' +finite→set-projective finite = + rec! (λ enum → Equiv→set-projective (enum e⁻¹) Fin-set-projective) + (Finite.enumeration finite) +``` + +## Taboos + +As it turns out, the finite types are the *only* types that are projective +constructively! The general sketch of the taboo is that it is consistent that: + +1. Propositions are projective. +2. Every infinite set admits an injection from `Nat`{.Agda}. In other + words, every infinite set is Dedekind-infinite. +3. Countable choice fails (IE: the natural numbers are not set-projective). + +The existence of such a model is out of scope for this page, so we will +focus our attention on the internal portion of the argument. In particular, +we will prove that if propositions are set-projective, then the existence of an +Dedekind-infinite set-projective type implies countable choice. + +First, note that if propositions are set-projective, then the [[image]] of +every function into a set-projective type is itself set-projective. This +follows directly from the definition of images, along with closure of +projectives under `Σ`{.Agda}. + +```agda +module _ + (props-projective : ∀ {ℓ ℓ'} → (A : Type ℓ) → is-prop A → is-set-projective A ℓ') + where + + props-projective→image-projective + : ∀ {A : Type ℓ} {B : Type ℓ'} + → (f : A → B) + → is-set-projective B (ℓ ⊔ ℓ' ⊔ ℓ'') + → is-set-projective (image f) ℓ'' + props-projective→image-projective f B-pro = + Σ-set-projective B-pro λ b → props-projective _ (hlevel 1) +``` + +This in turn implies that set-projective types are stable under embeddings, +as the image of an [[embedding]] $f : A \mono B$ is equivalent to $A$. + +```agda + props-projective+is-embedding→set-projective + : ∀ {A : Type ℓ} {B : Type ℓ'} {f : A → B} + → is-embedding f + → is-set-projective B (ℓ ⊔ ℓ' ⊔ ℓ'') + → is-set-projective A ℓ'' + props-projective+is-embedding→set-projective {f = f} f-emb B-pro = + Equiv→set-projective + (is-embedding→image-equiv f-emb e⁻¹) + (props-projective→image-projective f B-pro) +``` + +If we specialise this result to embeddings $Nat \mono A$, then we +obtain countable choice from the existence of a Dedekind-infinite type. + +```agda + props-projective+dedekind-infinite-projective→countable-choice + : ∀ {A : Type ℓ} {f : Nat → A} + → is-embedding f + → is-set-projective A (ℓ ⊔ ℓ') + → is-set-projective Nat ℓ' + props-projective+dedekind-infinite-projective→countable-choice = + props-projective+is-embedding→set-projective +``` + +Note that the set-projectivity of propositions is itself a taboo: in particular, +every proposition is set-projective if and only if every set has split support. +The following proof is adapted from [@Kraus-Escardó-Coquand-Altenkirch:2016]. + +We will start with the reverse direction. Suppose that every proposition +is set projective, and let $A$ be a set. The truncation of $A$ is a propositon, +and the constant family $\| A \| \to A$ is a set-indexed family, so projectivity +of $\| A \|$ directly gives us split support. + +```agda +props-projective→split-support + : ∀ {ℓ} + → ((A : Type ℓ) → is-prop A → is-set-projective A ℓ) + → ∀ (A : Type ℓ) → is-set A → ∥ (∥ A ∥ → A) ∥ +props-projective→split-support props-projective A A-set = + props-projective ∥ A ∥ (hlevel 1) (λ _ → A) (λ _ → A-set) id +``` + +For the forward direction, suppose that every set has split support, +let $A$ be a proposition, and $P : A \to \set$ a family of merely inhabited +sets. Note that the type $\Sigma A P$ is a set, so it must have split +support $s : \| \| \Sigma A P \| \to \Sigma A P \|$. Moreover, $P(a)$ +is always merely inhabited, so we can readily show that $\| A \to \Sigma A P\|$. +Finally, $A$ is a proposition, so we can obtain a $P(a)$ from $\Sigma A P$ +for any $a : A$; if we combine this with our previous observation, we +immediately get $\| (a : A) \to P(a) \|$. + +```agda +split-support→props-projective + : ∀ {ℓ} + → (∀ (A : Type ℓ) → is-set A → ∥ (∥ A ∥ → A) ∥) + → (A : Type ℓ) → is-prop A → is-set-projective A ℓ +split-support→props-projective split-support A A-prop P P-set ∥P∥ = do + s ← split-support (Σ[ a ∈ A ] P a) (Σ-is-hlevel 2 (is-prop→is-set A-prop) P-set) + pure λ a → subst P (A-prop _ _) (snd (s (∥-∥-map (λ p → (a , p)) (∥P∥ a)))) +``` diff --git a/src/Data/Set/Surjection.lagda.md b/src/Data/Set/Surjection.lagda.md index 7b60b2aba..695798863 100644 --- a/src/Data/Set/Surjection.lagda.md +++ b/src/Data/Set/Surjection.lagda.md @@ -2,6 +2,7 @@ ```agda open import Cat.Diagram.Coequaliser.RegularEpi open import Cat.Diagram.Coequaliser +open import Cat.Morphism.StrongEpi open import Cat.Prelude open import Homotopy.Connectedness @@ -33,7 +34,7 @@ together, we get _every epimorphism is regular_ as a corollary. [regular epimorphisms]: Cat.Diagram.Coequaliser.RegularEpi.html#regular-epimorphisms [epimorphism]: Cat.Morphism.html#epis -## Surjections are epic +## Surjections are epic {defines="surjections-are-epic"} This is the straightforward direction: We know (since $\Sets$ has pullbacks) that if a morphism is going to be a regular epimorphism, then @@ -81,7 +82,26 @@ surjectivity out of the way, we get what we wanted. (surj a) ``` -# Epis are surjective + + +# Epis are surjective {defines="epis-are-surjective"} Now _this_ is the hard direction. Our proof follows [@Rijke:2015, §2.9], so if the exposition below doesn't make a lot of sense, be sure to check diff --git a/src/bibliography.bibtex b/src/bibliography.bibtex index 5c6142112..456148fb1 100644 --- a/src/bibliography.bibtex +++ b/src/bibliography.bibtex @@ -265,3 +265,10 @@ primaryClass={math.CT}, url={https://arxiv.org/abs/1801.06488}, } + +@article{Kraus-Escardó-Coquand-Altenkirch:2016, + title={Notions of Anonymous Existence in Martin-Löf Type Theory}, + journal={Log. Methods Comput. Sci.}, author={Kraus, Nicolai and Escardó, M. and Coquand, T. and Altenkirch, Thorsten}, + year={2016}, + language={en} +} \ No newline at end of file From 7cc9bf7bbe90be5491e0d64da90a36afa29a540b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Am=C3=A9lia?= Date: Mon, 5 Aug 2024 09:02:45 -0300 Subject: [PATCH 3/9] wip: prime numbers (#416) numbers :-1: --- src/1Lab/Classical.lagda.md | 2 +- src/1Lab/Equiv.lagda.md | 8 +- src/1Lab/HLevel.lagda.md | 2 +- src/1Lab/Membership.agda | 2 +- src/1Lab/Path.lagda.md | 5 +- src/1Lab/Path/IdentitySystem.lagda.md | 11 +- src/1Lab/Resizing.lagda.md | 6 +- src/1Lab/Rewrite.agda | 12 - src/1Lab/Type.lagda.md | 2 +- src/1Lab/Underlying.agda | 5 + src/1Lab/Univalence.lagda.md | 2 +- src/Algebra/Magma.lagda.md | 2 +- src/Algebra/Ring/Solver.agda | 1 - src/Cat/Base.lagda.md | 13 +- src/Cat/CartesianClosed/Instances/PSh.agda | 4 +- src/Cat/Displayed/Comprehension.lagda.md | 2 +- src/Cat/Displayed/Total/Op.lagda.md | 17 +- src/Cat/Instances/Comma.lagda.md | 2 +- src/Cat/Instances/Shape/Involution.lagda.md | 2 +- src/Cat/Instances/Shape/Parallel.lagda.md | 6 +- src/Cat/Instances/Shape/Two.lagda.md | 4 +- src/Cat/Instances/StrictCat/Cohesive.lagda.md | 2 +- src/Cat/Internal/Opposite.lagda.md | 6 +- src/Cat/Monoidal/Instances/Day.lagda.md | 6 +- src/Data/Bool.lagda.md | 94 +----- src/Data/Bool/Base.lagda.md | 129 ++++++++ src/Data/Dec/Base.lagda.md | 8 +- src/Data/Fin/Base.lagda.md | 12 +- src/Data/Fin/Properties.lagda.md | 47 ++- src/Data/Id/Base.lagda.md | 11 +- src/Data/Int/HIT.lagda.md | 8 +- src/Data/List/Base.lagda.md | 6 +- src/Data/List/Properties.lagda.md | 1 - src/Data/List/Quantifiers.lagda.md | 19 +- src/Data/Nat/Base.lagda.md | 40 +-- src/Data/Nat/DivMod.lagda.md | 118 +++++-- src/Data/Nat/Divisible.lagda.md | 29 +- src/Data/Nat/Divisible/GCD.lagda.md | 41 +++ src/Data/Nat/Order.lagda.md | 129 ++++++-- src/Data/Nat/Prime.lagda.md | 291 ++++++++++++++++++ src/Data/Nat/Properties.lagda.md | 21 +- src/Data/Sum/Properties.lagda.md | 2 +- src/Homotopy/Space/Delooping.lagda.md | 5 +- .../Propositional/Classical/SAT.lagda.md | 6 +- src/Meta/Append.lagda.md | 2 +- src/Prim/Data/Nat.lagda.md | 4 +- src/Prim/Data/Sigma.lagda.md | 2 +- src/Prim/Kan.lagda.md | 4 +- 48 files changed, 860 insertions(+), 293 deletions(-) delete mode 100644 src/1Lab/Rewrite.agda create mode 100644 src/Data/Bool/Base.lagda.md create mode 100644 src/Data/Nat/Prime.lagda.md diff --git a/src/1Lab/Classical.lagda.md b/src/1Lab/Classical.lagda.md index 743cdfd4f..42220d040 100644 --- a/src/1Lab/Classical.lagda.md +++ b/src/1Lab/Classical.lagda.md @@ -153,7 +153,7 @@ gives us a section $\Sigma P \to 2$. ```agda module _ (split : Surjections-split) (P : Ω) where section : ∥ ((x : Susp ∣ P ∣) → fibre 2→Σ x) ∥ - section = split Bool-is-set (Susp-prop-is-set (hlevel 1)) 2→Σ 2→Σ-surjective + section = split (hlevel 2) (Susp-prop-is-set (hlevel 1)) 2→Σ 2→Σ-surjective ``` But a section is always injective, and the booleans are [[discrete]], so we can diff --git a/src/1Lab/Equiv.lagda.md b/src/1Lab/Equiv.lagda.md index 1a8ef473d..520346896 100644 --- a/src/1Lab/Equiv.lagda.md +++ b/src/1Lab/Equiv.lagda.md @@ -202,7 +202,7 @@ point with, rather than being a property of points. ```agda fibre : (A → B) → B → Type _ -fibre {A = A} f y = Σ[ x ∈ A ] (f x ≡ y) +fibre {A = A} f y = Σ[ x ∈ A ] f x ≡ y ``` ::: @@ -1094,13 +1094,13 @@ lift-inj p = ap Lift.lower p fibre-∘-≃ : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → {f : B → C} {g : A → B} - → ∀ c → fibre (f ∘ g) c ≃ (Σ[ (b , _) ∈ fibre f c ] (fibre g b)) + → ∀ c → fibre (f ∘ g) c ≃ (Σ[ (b , _) ∈ fibre f c ] fibre g b) fibre-∘-≃ {f = f} {g = g} c = Iso→Equiv (fwd , iso bwd invl invr) where - fwd : fibre (f ∘ g) c → Σ[ (b , _) ∈ fibre f c ] (fibre g b) + fwd : fibre (f ∘ g) c → Σ[ (b , _) ∈ fibre f c ] fibre g b fwd (a , p) = ((g a) , p) , (a , refl) - bwd : Σ[ (b , _) ∈ fibre f c ] (fibre g b) → fibre (f ∘ g) c + bwd : Σ[ (b , _) ∈ fibre f c ] fibre g b → fibre (f ∘ g) c bwd ((b , p) , (a , q)) = a , ap f q ∙ p invl : ∀ x → fwd (bwd x) ≡ x diff --git a/src/1Lab/HLevel.lagda.md b/src/1Lab/HLevel.lagda.md index 6d789e1f8..a6a6c2c88 100644 --- a/src/1Lab/HLevel.lagda.md +++ b/src/1Lab/HLevel.lagda.md @@ -204,7 +204,7 @@ proof (used in the *definition* of path induction), we can demonstrate the application of path induction: ```agda -_ : (a : A) → is-contr (Σ[ b ∈ A ] (b ≡ a)) +_ : (a : A) → is-contr (Σ[ b ∈ A ] b ≡ a) _ = λ t → record { centre = (t , refl) ; paths = λ (s , p) → J' (λ s t p → (t , refl) ≡ (s , p)) (λ t → refl) p diff --git a/src/1Lab/Membership.agda b/src/1Lab/Membership.agda index 60f1df61a..7fedf95f5 100644 --- a/src/1Lab/Membership.agda +++ b/src/1Lab/Membership.agda @@ -45,7 +45,7 @@ infix 30 _∉_ ∫ₚ : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {ℙX : Type ℓ'} ⦃ m : Membership X ℙX ℓ'' ⦄ → ℙX → Type _ -∫ₚ {X = X} P = Σ[ x ∈ X ] (x ∈ P) +∫ₚ {X = X} P = Σ[ x ∈ X ] x ∈ P -- Notation typeclass for _⊆_. We could always define -- diff --git a/src/1Lab/Path.lagda.md b/src/1Lab/Path.lagda.md index 3ccdd94ff..c0f9fb7c9 100644 --- a/src/1Lab/Path.lagda.md +++ b/src/1Lab/Path.lagda.md @@ -1044,7 +1044,7 @@ inhabitant! ```agda Singleton : ∀ {ℓ} {A : Type ℓ} → A → Type _ -Singleton x = Σ[ y ∈ _ ] (x ≡ y) +Singleton x = Σ[ y ∈ _ ] x ≡ y ``` We're given an inhabitant $y : A$ and a path $p : x \is y$. To identify @@ -2378,5 +2378,8 @@ _$ₚ_ : ∀ {ℓ₁ ℓ₂} {A : Type ℓ₁} {B : A → Type ℓ₂} {f g : → f ≡ g → ∀ x → f x ≡ g x (f $ₚ x) i = f i x {-# INLINE _$ₚ_ #-} + +_≠_ : ∀ {ℓ} {A : Type ℓ} → A → A → Type ℓ +x ≠ y = ¬ x ≡ y ``` --> diff --git a/src/1Lab/Path/IdentitySystem.lagda.md b/src/1Lab/Path/IdentitySystem.lagda.md index a2cc7a102..97f0e82f7 100644 --- a/src/1Lab/Path/IdentitySystem.lagda.md +++ b/src/1Lab/Path/IdentitySystem.lagda.md @@ -332,12 +332,7 @@ This is known as **Hedberg's theorem**. opaque Discrete→is-set : ∀ {ℓ} {A : Type ℓ} → Discrete A → is-set A - Discrete→is-set {A = A} dec = - identity-system→hlevel 1 (¬¬-stable-identity-system stable) λ x y f g → - funext λ h → absurd (g h) - where - stable : {x y : A} → ¬ ¬ x ≡ y → x ≡ y - stable {x = x} {y = y} ¬¬p with dec {x} {y} - ... | yes p = p - ... | no ¬p = absurd (¬¬p ¬p) + Discrete→is-set {A = A} dec = identity-system→hlevel 1 + (¬¬-stable-identity-system (dec→dne ⦃ dec ⦄)) + λ x y f g → funext λ h → absurd (g h) ``` diff --git a/src/1Lab/Resizing.lagda.md b/src/1Lab/Resizing.lagda.md index 4555c4e39..56bb6f318 100644 --- a/src/1Lab/Resizing.lagda.md +++ b/src/1Lab/Resizing.lagda.md @@ -227,9 +227,9 @@ disjunctions, and implications. diff --git a/src/1Lab/Rewrite.agda b/src/1Lab/Rewrite.agda deleted file mode 100644 index d2ad83d6f..000000000 --- a/src/1Lab/Rewrite.agda +++ /dev/null @@ -1,12 +0,0 @@ -open import 1Lab.Path -open import 1Lab.Type - -module 1Lab.Rewrite where - -data _≡rw_ {ℓ} {A : Type ℓ} (x : A) : A → SSet ℓ where - idrw : x ≡rw x - -{-# BUILTIN REWRITE _≡rw_ #-} - -postulate - make-rewrite : ∀ {ℓ} {A : Type ℓ} {x y : A} → Path A x y → x ≡rw y diff --git a/src/1Lab/Type.lagda.md b/src/1Lab/Type.lagda.md index 559cae942..6f29383f9 100644 --- a/src/1Lab/Type.lagda.md +++ b/src/1Lab/Type.lagda.md @@ -49,7 +49,7 @@ absurd () ¬_ : ∀ {ℓ} → Type ℓ → Type ℓ ¬ A = A → ⊥ -infix 3 ¬_ +infix 6 ¬_ ``` :::{.definition #product-type} diff --git a/src/1Lab/Underlying.agda b/src/1Lab/Underlying.agda index 9c90c96f0..068c8ed12 100644 --- a/src/1Lab/Underlying.agda +++ b/src/1Lab/Underlying.agda @@ -6,6 +6,8 @@ open import 1Lab.HLevel open import 1Lab.Path open import 1Lab.Type hiding (Σ-syntax) +open import Data.Bool.Base + module 1Lab.Underlying where -- Notation class for types which have a chosen projection into a @@ -54,6 +56,9 @@ instance Underlying-Lift ⦃ ua ⦄ .ℓ-underlying = ua .ℓ-underlying Underlying-Lift .⌞_⌟ x = ⌞ x .Lift.lower ⌟ + Underlying-Bool : Underlying Bool + Underlying-Bool = record { ⌞_⌟ = So } + -- The converse of to-is-true, generalised slightly. If P and Q are -- identical inhabitants of some type of structured types, and Q's -- underlying type is contractible, then P is inhabited. P's underlying diff --git a/src/1Lab/Univalence.lagda.md b/src/1Lab/Univalence.lagda.md index 427d8b624..a5d29b59b 100644 --- a/src/1Lab/Univalence.lagda.md +++ b/src/1Lab/Univalence.lagda.md @@ -95,7 +95,7 @@ open import 1Lab.Equiv.FromPath ```agda Glue : (A : Type ℓ) → {φ : I} - → (Te : Partial φ (Σ[ T ∈ Type ℓ' ] (T ≃ A))) + → (Te : Partial φ (Σ[ T ∈ Type ℓ' ] T ≃ A)) → Type ℓ' ``` diff --git a/src/Algebra/Magma.lagda.md b/src/Algebra/Magma.lagda.md index 3a84170f9..c7a0e0cf3 100644 --- a/src/Algebra/Magma.lagda.md +++ b/src/Algebra/Magma.lagda.md @@ -138,7 +138,7 @@ set, this obviously defines a magma: ```agda private is-magma-imp : is-magma imp - is-magma-imp .has-is-set = Bool-is-set + is-magma-imp .has-is-set = hlevel 2 ``` We show it is not commutative or associative by giving counterexamples: diff --git a/src/Algebra/Ring/Solver.agda b/src/Algebra/Ring/Solver.agda index ebe85fd38..bf1107915 100644 --- a/src/Algebra/Ring/Solver.agda +++ b/src/Algebra/Ring/Solver.agda @@ -12,7 +12,6 @@ Horner normal forms are not sparse). open import 1Lab.Reflection.Variables open import 1Lab.Reflection.Solver open import 1Lab.Reflection -open import 1Lab.Rewrite open import Algebra.Ring.Cat.Initial open import Algebra.Ring.Commutative diff --git a/src/Cat/Base.lagda.md b/src/Cat/Base.lagda.md index e21d71a83..34292d46e 100644 --- a/src/Cat/Base.lagda.md +++ b/src/Cat/Base.lagda.md @@ -8,7 +8,6 @@ open import 1Lab.Extensionality open import 1Lab.HLevel.Closure open import 1Lab.Reflection open import 1Lab.Underlying -open import 1Lab.Rewrite open import 1Lab.HLevel open import 1Lab.Equiv open import 1Lab.Path @@ -198,10 +197,7 @@ C^op^op≡C {C = C} i = precat i where @@ -312,12 +308,7 @@ F^op^op≡F {F = F} i .Functor.F₁ = F .Functor.F₁ F^op^op≡F {F = F} i .Functor.F-id = F .Functor.F-id F^op^op≡F {F = F} i .Functor.F-∘ = F .Functor.F-∘ -private - functor-double-dual - : ∀ {o ℓ o' ℓ'} {C : Precategory o ℓ} {D : Precategory o' ℓ'} {F : Functor C D} - → Functor.op (Functor.op F) ≡rw F - functor-double-dual = make-rewrite F^op^op≡F -{-# REWRITE functor-double-dual #-} +{-# REWRITE F^op^op≡F #-} ``` --> diff --git a/src/Cat/CartesianClosed/Instances/PSh.agda b/src/Cat/CartesianClosed/Instances/PSh.agda index ce2fab861..ebf830d35 100644 --- a/src/Cat/CartesianClosed/Instances/PSh.agda +++ b/src/Cat/CartesianClosed/Instances/PSh.agda @@ -53,7 +53,7 @@ module _ {o ℓ κ} {C : Precategory o ℓ} where open is-pullback pb-path - : ∀ {i} {x y : Σ[ x ∈ X.₀ i ] Σ[ y ∈ Y.₀ i ] (f.η i x ≡ g.η i y)} + : ∀ {i} {x y : Σ[ x ∈ X.₀ i ] Σ[ y ∈ Y.₀ i ] f.η i x ≡ g.η i y} → x .fst ≡ y .fst → x .snd .fst ≡ y .snd .fst → x ≡ y @@ -65,7 +65,7 @@ module _ {o ℓ κ} {C : Precategory o ℓ} where i j pb : Pullback (PSh κ C) f g - pb .apex .F₀ i = el! (Σ[ x ∈ X.₀ i ] Σ[ y ∈ Y.₀ i ] (f.η i x ≡ g.η i y)) + pb .apex .F₀ i = el! (Σ[ x ∈ X.₀ i ] Σ[ y ∈ Y.₀ i ] f.η i x ≡ g.η i y) pb .apex .F₁ {x} {y} h (a , b , p) = X.₁ h a , Y.₁ h b , path where abstract path : f.η y (X.₁ h a) ≡ g.η y (Y.₁ h b) path = happly (f.is-natural _ _ _) _ diff --git a/src/Cat/Displayed/Comprehension.lagda.md b/src/Cat/Displayed/Comprehension.lagda.md index 251a1c359..cc09e56ee 100644 --- a/src/Cat/Displayed/Comprehension.lagda.md +++ b/src/Cat/Displayed/Comprehension.lagda.md @@ -164,7 +164,7 @@ with projections. _⨾ˢ_ : ∀ {Γ Δ x y} (σ : Hom Γ Δ) → Hom[ σ ] x y → Hom (Γ ⨾ x) (Δ ⨾ y) σ ⨾ˢ f = F₁' f .to - infixl 5 _⨾ˢ_ + infixl 8 _⨾ˢ_ sub-proj : ∀ {Γ Δ x y} {σ : Hom Γ Δ} → (f : Hom[ σ ] x y) → πᶜ ∘ (σ ⨾ˢ f) ≡ σ ∘ πᶜ sub-proj f = sym $ F₁' f .commute diff --git a/src/Cat/Displayed/Total/Op.lagda.md b/src/Cat/Displayed/Total/Op.lagda.md index c8c69c3c2..4f49c9d7d 100644 --- a/src/Cat/Displayed/Total/Op.lagda.md +++ b/src/Cat/Displayed/Total/Op.lagda.md @@ -1,7 +1,5 @@ @@ -180,12 +173,6 @@ fibre-functor-total-op-total-op {ℰ = ℰ} {y = y} {F = F} i .F-∘ f g = i where open Precategory (Fibre ℰ y) -private - fibre-functor-double-dual - : ∀ {o ℓ o' ℓ'} {ℬ : Precategory o ℓ} {ℰ : Displayed ℬ o' ℓ'} {x y} - → {F : Functor (Fibre ℰ x) (Fibre ℰ y)} - → fibre-functor-total-op (fibre-functor-total-op F) ≡rw F - fibre-functor-double-dual = make-rewrite fibre-functor-total-op-total-op -{-# REWRITE fibre-functor-double-dual #-} +{-# REWRITE fibre-functor-total-op-total-op #-} ``` --> diff --git a/src/Cat/Instances/Comma.lagda.md b/src/Cat/Instances/Comma.lagda.md index 45e20d598..009cdc52a 100644 --- a/src/Cat/Instances/Comma.lagda.md +++ b/src/Cat/Instances/Comma.lagda.md @@ -248,7 +248,7 @@ module _ {A : Precategory ao ah} {B : Precategory bo bh} where open ↓Obj open ↓Hom - infix 5 _↙_ _↘_ + infix 8 _↙_ _↘_ _↙_ : A.Ob → Functor B A → Precategory _ _ X ↙ T = !Const X ↓ T diff --git a/src/Cat/Instances/Shape/Involution.lagda.md b/src/Cat/Instances/Shape/Involution.lagda.md index 6bd914ee8..a9b5ee85e 100644 --- a/src/Cat/Instances/Shape/Involution.lagda.md +++ b/src/Cat/Instances/Shape/Involution.lagda.md @@ -24,7 +24,7 @@ composition operation. ∙⤮∙ : Precategory lzero lzero ∙⤮∙ .Precategory.Ob = ⊤ ∙⤮∙ .Precategory.Hom _ _ = Bool -∙⤮∙ .Precategory.Hom-set _ _ = Bool-is-set +∙⤮∙ .Precategory.Hom-set _ _ = hlevel 2 ∙⤮∙ .Precategory.id = false ∙⤮∙ .Precategory._∘_ = xor ∙⤮∙ .Precategory.idr f = xor-falser f diff --git a/src/Cat/Instances/Shape/Parallel.lagda.md b/src/Cat/Instances/Shape/Parallel.lagda.md index 8d962b429..23eb71651 100644 --- a/src/Cat/Instances/Shape/Parallel.lagda.md +++ b/src/Cat/Instances/Shape/Parallel.lagda.md @@ -39,9 +39,9 @@ parallel arrows between them. It is the shape of [[equaliser]] and diff --git a/src/Cat/Monoidal/Instances/Day.lagda.md b/src/Cat/Monoidal/Instances/Day.lagda.md index ab05a00f7..5976d5bbc 100644 --- a/src/Cat/Monoidal/Instances/Day.lagda.md +++ b/src/Cat/Monoidal/Instances/Day.lagda.md @@ -1,7 +1,5 @@ diff --git a/src/Data/Bool/Base.lagda.md b/src/Data/Bool/Base.lagda.md new file mode 100644 index 000000000..8af287c3c --- /dev/null +++ b/src/Data/Bool/Base.lagda.md @@ -0,0 +1,129 @@ + + +```agda +module Data.Bool.Base where +``` + + + +```agda +true≠false : ¬ true ≡ false +true≠false p = subst P p tt where + P : Bool → Type + P false = ⊥ + P true = ⊤ +``` + + + +```agda +not : Bool → Bool +not true = false +not false = true + +and or : Bool → Bool → Bool +and false y = false +and true y = y + +or false y = y +or true y = true +``` + +```agda +instance + Discrete-Bool : Discrete Bool + Discrete-Bool {false} {false} = yes refl + Discrete-Bool {false} {true} = no false≠true + Discrete-Bool {true} {false} = no true≠false + Discrete-Bool {true} {true} = yes refl +``` + + + +```agda +x≠true→x≡false : (x : Bool) → x ≠ true → x ≡ false +x≠true→x≡false false p = refl +x≠true→x≡false true p = absurd (p refl) + +x≠false→x≡true : (x : Bool) → x ≠ false → x ≡ true +x≠false→x≡true false p = absurd (p refl) +x≠false→x≡true true p = refl +``` + +```agda +is-true : Bool → Type +is-true true = ⊤ +is-true false = ⊥ + +record So (b : Bool) : Type where + field + is-so : is-true b + +pattern oh = record { is-so = tt } +``` + + + + diff --git a/src/Data/Dec/Base.lagda.md b/src/Data/Dec/Base.lagda.md index 4ccd371bf..79b0b0b8a 100644 --- a/src/Data/Dec/Base.lagda.md +++ b/src/Data/Dec/Base.lagda.md @@ -47,9 +47,11 @@ Dec-rec = Dec-elim _ ```agda recover : ∀ {ℓ} {A : Type ℓ} ⦃ d : Dec A ⦄ → .A → A recover ⦃ yes x ⦄ _ = x -recover {A = A} ⦃ no ¬x ⦄ x = go (¬x x) where - go : .⊥ → A - go () +recover ⦃ no ¬x ⦄ x = absurd (¬x x) + +dec→dne : ∀ {ℓ} {A : Type ℓ} ⦃ d : Dec A ⦄ → ¬ ¬ A → A +dec→dne ⦃ yes x ⦄ _ = x +dec→dne ⦃ no ¬x ⦄ ¬¬x = absurd (¬¬x ¬x) ``` --> diff --git a/src/Data/Fin/Base.lagda.md b/src/Data/Fin/Base.lagda.md index d1c24c44b..2efc4e27a 100644 --- a/src/Data/Fin/Base.lagda.md +++ b/src/Data/Fin/Base.lagda.md @@ -187,6 +187,10 @@ Fin-elim Fin-elim P pfzero pfsuc fzero = pfzero Fin-elim P pfzero pfsuc (fsuc x) = pfsuc x (Fin-elim P pfzero pfsuc x) +fin-cons : ∀ {ℓ} {n} {P : Fin (suc n) → Type ℓ} → P 0 → (∀ x → P (fsuc x)) → ∀ x → P x +fin-cons p0 ps fzero = p0 +fin-cons p0 ps (fsuc x) = ps x + fin-absurd : Fin 0 → ⊥ fin-absurd () ``` @@ -203,11 +207,11 @@ sets. ```agda _≤_ : ∀ {n} → Fin n → Fin n → Type i ≤ j = to-nat i Nat.≤ to-nat j -infix 3 _≤_ +infix 7 _≤_ _<_ : ∀ {n} → Fin n → Fin n → Type i < j = to-nat i Nat.< to-nat j -infix 3 _<_ +infix 7 _<_ ``` Next, we define a pair of functions `squish`{.Agda} and `skip`{.Agda}, @@ -240,7 +244,7 @@ $n$ values of $\bb{N}$. ```agda ℕ< : Nat → Type -ℕ< x = Σ[ n ∈ Nat ] (n Nat.< x) +ℕ< x = Σ[ n ∈ Nat ] n Nat.< x from-ℕ< : ∀ {n} → ℕ< n → Fin n from-ℕ< {n = suc n} (zero , q) = fzero @@ -271,7 +275,7 @@ split-+ {m = zero} i = inr i split-+ {m = suc m} fzero = inl fzero split-+ {m = suc m} (fsuc i) = ⊎-map fsuc id (split-+ i) -avoid : ∀ {n} (i j : Fin (suc n)) → (¬ i ≡ j) → Fin n +avoid : ∀ {n} (i j : Fin (suc n)) → i ≠ j → Fin n avoid fzero fzero i≠j = absurd (i≠j refl) avoid {n = suc n} fzero (fsuc j) i≠j = j avoid {n = suc n} (fsuc i) fzero i≠j = fzero diff --git a/src/Data/Fin/Properties.lagda.md b/src/Data/Fin/Properties.lagda.md index 1c6da5354..e9e37fed9 100644 --- a/src/Data/Fin/Properties.lagda.md +++ b/src/Data/Fin/Properties.lagda.md @@ -3,6 +3,7 @@ open import 1Lab.Prelude open import Data.Fin.Base +open import Data.Nat.Base using (s≤s) open import Data.Dec open import Data.Sum @@ -167,7 +168,7 @@ Fin≃ℕ< : ∀ {n} → Fin n ≃ ℕ< n Fin≃ℕ< {n} = to-ℕ< , is-iso→is-equiv (iso from-ℕ< (to-from-ℕ< {n}) from-to-ℕ<) avoid-injective - : ∀ {n} (i : Fin (suc n)) {j k : Fin (suc n)} {i≠j : ¬ i ≡ j} {i≠k : ¬ i ≡ k} + : ∀ {n} (i : Fin (suc n)) {j k : Fin (suc n)} {i≠j : i ≠ j} {i≠k : i ≠ k} → avoid i j i≠j ≡ avoid i k i≠k → j ≡ k avoid-injective fzero {fzero} {k} {i≠j} p = absurd (i≠j refl) avoid-injective fzero {fsuc j} {fzero} {i≠k = i≠k} p = absurd (i≠k refl) @@ -191,8 +192,7 @@ Fin-suc-Π Fin-suc-Π = Iso→Equiv λ where .fst f → f fzero , (λ x → f (fsuc x)) - .snd .is-iso.inv (z , f) fzero → z - .snd .is-iso.inv (z , f) (fsuc x) → f x + .snd .is-iso.inv (z , s) → fin-cons z s .snd .is-iso.rinv x → refl @@ -287,17 +287,48 @@ $\Pi$-types and $\Sigma$-types, respectively. ```agda instance - Fin-exhaustible + Dec-Fin-∀ : ∀ {n ℓ} {A : Fin n → Type ℓ} → ⦃ ∀ {x} → Dec (A x) ⦄ → Dec (∀ x → A x) - Fin-exhaustible {n} ⦃ d ⦄ = Fin-Monoidal n λ _ → d + Dec-Fin-∀ {n} ⦃ d ⦄ = Fin-Monoidal n (λ _ → d) - Fin-omniscient + Dec-Fin-Σ : ∀ {n ℓ} {A : Fin n → Type ℓ} → ⦃ ∀ {x} → Dec (A x) ⦄ → Dec (Σ (Fin n) A) - Fin-omniscient {n} ⦃ d ⦄ = Fin-Alternative n λ _ → d + Dec-Fin-Σ {n} ⦃ d ⦄ = Fin-Alternative n λ _ → d ``` +```agda +Fin-omniscience + : ∀ {n ℓ} (P : Fin n → Type ℓ) ⦃ _ : ∀ {x} → Dec (P x) ⦄ + → (Σ[ j ∈ Fin n ] P j × ∀ k → P k → j ≤ k) ⊎ (∀ x → ¬ P x) +Fin-omniscience {zero} P = inr λ () +Fin-omniscience {suc n} P with holds? (P 0) +... | yes here = inl (0 , here , λ _ _ → 0≤x) +... | no ¬here with Fin-omniscience (P ∘ fsuc) +... | inl (ix , pix , least) = inl (fsuc ix , pix , fin-cons (λ here → absurd (¬here here)) λ i pi → Nat.s≤s (least i pi)) +... | inr nowhere = inr (fin-cons ¬here nowhere) +``` + + + ## Injections and surjections The standard finite sets are **Dedekind-finite**, which means that every @@ -347,7 +378,7 @@ avoid-insert → (ρ : Fin n → A) → (i : Fin (suc n)) (a : A) → (j : Fin (suc n)) - → (i≠j : ¬ i ≡ j) + → (i≠j : i ≠ j) → (ρ [ i ≔ a ]) j ≡ ρ (avoid i j i≠j) avoid-insert {n = n} ρ fzero a fzero i≠j = absurd (i≠j refl) avoid-insert {n = suc n} ρ fzero a (fsuc j) i≠j = refl diff --git a/src/Data/Id/Base.lagda.md b/src/Data/Id/Base.lagda.md index 0959b0190..d04d442a3 100644 --- a/src/Data/Id/Base.lagda.md +++ b/src/Data/Id/Base.lagda.md @@ -6,7 +6,6 @@ open import 1Lab.Path.IdentitySystem open import 1Lab.HLevel.Closure open import 1Lab.Type.Sigma open import 1Lab.Univalence -open import 1Lab.Rewrite open import 1Lab.HLevel open import 1Lab.Equiv open import 1Lab.Path @@ -114,13 +113,13 @@ opaque _≡ᵢ?_ : ∀ {ℓ} {A : Type ℓ} ⦃ _ : Discrete A ⦄ (x y : A) → Dec (x ≡ᵢ y) x ≡ᵢ? y = discrete-id (x ≡? y) - ≡ᵢ?-default : ∀ {ℓ} {A : Type ℓ} {x y : A} {d : Discrete A} → (_≡ᵢ?_ ⦃ d ⦄ x y) ≡rw discrete-id d - ≡ᵢ?-default = make-rewrite refl + ≡ᵢ?-default : ∀ {ℓ} {A : Type ℓ} {x y : A} {d : Discrete A} → (_≡ᵢ?_ ⦃ d ⦄ x y) ≡ discrete-id d + ≡ᵢ?-default = refl - ≡ᵢ?-yes : ∀ {ℓ} {A : Type ℓ} {x : A} {d : Discrete A} → (_≡ᵢ?_ ⦃ d ⦄ x x) ≡rw yes reflᵢ - ≡ᵢ?-yes {d = d} = make-rewrite (case d return (λ d → discrete-id d ≡ yes reflᵢ) of λ where + ≡ᵢ?-yes : ∀ {ℓ} {A : Type ℓ} {x : A} {d : Discrete A} → (_≡ᵢ?_ ⦃ d ⦄ x x) ≡ yes reflᵢ + ≡ᵢ?-yes {d = d} = case d return (λ d → discrete-id d ≡ yes reflᵢ) of λ where (yes a) → ap yes (is-set→is-setᵢ (Discrete→is-set d) _ _ _ _) - (no ¬a) → absurd (¬a refl)) + (no ¬a) → absurd (¬a refl) {-# REWRITE ≡ᵢ?-default ≡ᵢ?-yes #-} diff --git a/src/Data/Int/HIT.lagda.md b/src/Data/Int/HIT.lagda.md index 02df4c096..49b2efed7 100644 --- a/src/Data/Int/HIT.lagda.md +++ b/src/Data/Int/HIT.lagda.md @@ -1,7 +1,6 @@ diff --git a/src/Data/Nat/Base.lagda.md b/src/Data/Nat/Base.lagda.md index 092851e38..ae1de2e0f 100644 --- a/src/Data/Nat/Base.lagda.md +++ b/src/Data/Nat/Base.lagda.md @@ -6,8 +6,8 @@ open import 1Lab.HLevel open import 1Lab.Path open import 1Lab.Type +open import Data.Bool.Base open import Data.Dec.Base -open import Data.Bool ``` --> @@ -104,23 +104,28 @@ private module _ where private @@ -181,7 +186,7 @@ _^_ : Nat → Nat → Nat x ^ zero = 1 x ^ suc y = x * (x ^ y) -infixr 8 _^_ +infixr 10 _^_ ``` ## Ordering @@ -215,8 +220,7 @@ instance {-# INCOHERENT x≤x x≤sucy #-} Positive : Nat → Type -Positive zero = ⊥ -Positive (suc n) = ⊤ +Positive n = 1 ≤ n ``` --> @@ -226,7 +230,7 @@ re-using the definition of `_≤_`{.Agda}. ```agda _<_ : Nat → Nat → Type m < n = suc m ≤ n -infix 3 _<_ _≤_ +infix 7 _<_ _≤_ ``` As an "ordering combinator", we can define the _maximum_ of two natural diff --git a/src/Data/Nat/DivMod.lagda.md b/src/Data/Nat/DivMod.lagda.md index 08e03d0fe..6009ce464 100644 --- a/src/Data/Nat/DivMod.lagda.md +++ b/src/Data/Nat/DivMod.lagda.md @@ -46,9 +46,18 @@ The easy case is to divide zero by anything, in which case the result is zero with remainder zero. The more interesting case comes when we have some successor, and we want to divide it. + + +```agda + divide-pos : ∀ a b → .⦃ _ : Positive b ⦄ → DivMod a b + divide-pos zero (suc b) = divmod 0 0 refl (s≤s 0≤x) ``` It suffices to assume --- since $a$ is smaller than $1+a$ --- that we @@ -57,8 +66,8 @@ ordering on $\NN$ is trichotomous, we can proceed by cases on whether $(1 + r') < b$. ```agda -divide-pos (suc a) b with divide-pos a b -divide-pos (suc a) b | divmod q' r' p s with ≤-split (suc r') b + divide-pos (suc a) b with divide-pos a b + divide-pos (suc a) b | divmod q' r' p s with ≤-split (suc r') b ``` First, suppose that $1 + r' < b$, i.e., $1 + r'$ _can_ serve as a @@ -70,8 +79,8 @@ $$ $$ ```agda -divide-pos (suc a) b | divmod q' r' p s | inl r'+1 + ```agda instance Dec-∣ : ∀ {n m} → Dec (n ∣ m) @@ -138,5 +196,5 @@ instance let n∣r : (q' - q) * n ≡ r n∣r = monus-distribr q' q n ∙ sym (monus-swapl _ _ _ (sym (p ∙ recover α))) - in <-≤-asym (recover β) (m∣sn→m≤sn (q' - q , n∣r)) + in <-≤-asym (recover β) (m∣sn→m≤sn (q' - q , recover n∣r)) ``` diff --git a/src/Data/Nat/Divisible.lagda.md b/src/Data/Nat/Divisible.lagda.md index 1408a1a57..b6fefa833 100644 --- a/src/Data/Nat/Divisible.lagda.md +++ b/src/Data/Nat/Divisible.lagda.md @@ -34,7 +34,7 @@ _∣_ : Nat → Nat → Type zero ∣ y = y ≡ zero suc x ∣ y = fibre (_* suc x) y -infix 5 _∣_ +infix 7 _∣_ ``` In this way, we break the pathological case of $0 | 0$ by _decreeing_ it @@ -117,6 +117,14 @@ m∣sn→m≤sn : ∀ {x y} → x ∣ suc y → x ≤ suc y m∣sn→m≤sn {x} {y} p with ∣→fibre p ... | zero , p = absurd (zero≠suc p) ... | suc k , p = difference→≤ (k * x) p + +m∣n→m≤n : ∀ {m n} .⦃ _ : Positive n ⦄ → m ∣ n → m ≤ n +m∣n→m≤n {n = suc _} = m∣sn→m≤sn + +proper-divisor-< : ∀ {m n} .⦃ _ : Positive n ⦄ → m ≠ n → m ∣ n → m < n +proper-divisor-< m≠n m∣n with ≤-strengthen (m∣n→m≤n m∣n) +... | inl here = absurd (m≠n here) +... | inr there = there ``` This will let us establish the antisymmetry we were looking for: @@ -139,6 +147,9 @@ expect a number to divide its multiples. Fortunately, this is the case: ∣-*r : ∀ {x y} → x ∣ y * x ∣-*r {y = y} = fibre→∣ (y , refl) + +|-*l-pres : ∀ {n a b} → n ∣ b → n ∣ a * b +|-*l-pres {n} {a} {b} p1 with (q , α) ← ∣→fibre p1 = fibre→∣ (a * q , *-associative a q n ∙ ap (a *_) α) ``` If two numbers are multiples of $k$, then so is their sum. @@ -147,6 +158,22 @@ If two numbers are multiples of $k$, then so is their sum. ∣-+ : ∀ {k n m} → k ∣ n → k ∣ m → k ∣ (n + m) ∣-+ {zero} {n} {m} p q = ap (_+ m) p ∙ q ∣-+ {suc k} (x , p) (y , q) = x + y , *-distrib-+r x y (suc k) ∙ ap₂ _+_ p q + +∣-+-cancel : ∀ {n a b} → n ∣ a → n ∣ a + b → n ∣ b +∣-+-cancel {n} {a} {b} p1 p2 with (q , α) ← ∣→fibre p1 | (r , β) ← ∣→fibre p2 = fibre→∣ + (r - q , monus-distribr r q n ∙ ap₂ _-_ β α ∙ ap ((a + b) -_) (sym (+-zeror a)) ∙ monus-cancell a b 0) + +∣-+-cancel' : ∀ {n a b} → n ∣ b → n ∣ a + b → n ∣ a +∣-+-cancel' {n} {a} {b} p1 p2 = ∣-+-cancel {n} {b} {a} p1 (subst (n ∣_) (+-commutative a b) p2) +``` + +The only number that divides 1 is 1 itself: + +```agda +∣-1 : ∀ {n} → n ∣ 1 → n ≡ 1 +∣-1 {0} p = sym p +∣-1 {1} p = refl +∣-1 {suc (suc n)} (k , p) = *-is-oner k (2 + n) p ``` ## Even and odd numbers diff --git a/src/Data/Nat/Divisible/GCD.lagda.md b/src/Data/Nat/Divisible/GCD.lagda.md index 8309217f9..5eb910793 100644 --- a/src/Data/Nat/Divisible/GCD.lagda.md +++ b/src/Data/Nat/Divisible/GCD.lagda.md @@ -8,6 +8,7 @@ open import Data.Nat.Properties open import Data.Nat.Divisible open import Data.Nat.DivMod open import Data.Nat.Order +open import Data.Dec.Base open import Data.Nat.Base open import Data.Sum.Base ``` @@ -191,3 +192,43 @@ is-gcd-graphs-gcd {m = m} {n} {d} = prop-ext! (λ x → ap fst $ GCD-is-prop (gcd m n , Euclid.euclid m n .snd) (d , x)) (λ p → subst (is-gcd m n) p (Euclid.euclid m n .snd)) ``` + +## Euclid's lemma + +```agda +|-*-coprime-cancel + : ∀ n a b .⦃ _ : Positive n ⦄ .⦃ _ : Positive a ⦄ + → n ∣ a * b → is-gcd n a 1 → n ∣ b +|-*-coprime-cancel n a b div coprime = done where + E : Nat → Prop lzero + E x = el! (0 < x × n ∣ x * b) + + has : Σ[ x ∈ Nat ] ((x ∈ E) × (∀ y → (0 < y) × (n ∣ y * b) → x ≤ y)) + has = ℕ-well-ordered {P = E} (λ _ → auto) (inc (a , recover auto , div)) + + instance + _ : Positive (has .fst) + _ = has .snd .fst .fst + + step : ∀ x → x ∈ E → has .fst ∣ x + step x (xe , d) with divmod q r α β ← divide-pos x (has .fst) = + let + d' : n ∣ (q * has .fst * b + r * b) + d' = subst (n ∣_) (ap (_* b) (recover α) ∙ *-distrib-+r (q * has .fst) r b) d + + d'' : n ∣ r * b + d'' = ∣-+-cancel {n} {q * has .fst * b} {r * b} (subst (n ∣_) (sym (*-associative q (has .fst) b)) (|-*l-pres {a = q} (has .snd .fst .snd))) d' + in case r ≡? 0 of λ where + (yes p) → fibre→∣ (q , sym (recover α ∙ ap (q * has .fst +_) p ∙ +-zeror (q * has .fst))) + (no ¬p) → absurd (<-irrefl + (≤-antisym (≤-trans ≤-ascend (recover β)) (has .snd .snd r (nonzero→positive ¬p , d''))) + (recover β)) + + almost : has .fst ≡ 1 + almost = ∣-1 $ coprime .greatest {has .fst} + (step n (recover auto , ∣-*l)) + (step a (recover auto , div)) + + done : n ∣ b + done = subst (n ∣_) (ap (_* b) almost ∙ *-onel b) (has .snd .fst .snd) +``` diff --git a/src/Data/Nat/Order.lagda.md b/src/Data/Nat/Order.lagda.md index f47d58b8a..5d4186564 100644 --- a/src/Data/Nat/Order.lagda.md +++ b/src/Data/Nat/Order.lagda.md @@ -2,9 +2,12 @@ ```agda open import 1Lab.Prelude +open import Data.Bool.Base open import Data.Dec.Base open import Data.Nat.Base open import Data.Sum + +import Prim.Data.Nat as Prim ``` --> @@ -28,7 +31,19 @@ naturals automatically. ≤-refl : ∀ {x : Nat} → x ≤ x ≤-refl {zero} = 0≤x ≤-refl {suc x} = s≤s ≤-refl +``` + + +```agda ≤-trans : ∀ {x y z} → x ≤ y → y ≤ z → x ≤ z ≤-trans 0≤x 0≤x = 0≤x ≤-trans 0≤x (s≤s q) = 0≤x @@ -60,22 +75,73 @@ equivalence between $x \le y$ and $(1 + x) \le (1 + y)$. +### Properties of the strict order + +```agda +<-≤-asym : ∀ {x y} → x < y → ¬ (y ≤ x) +<-≤-asym {.(suc _)} {.(suc _)} (s≤s p) (s≤s q) = <-≤-asym p q + +<-asym : ∀ {x y} → x < y → ¬ (y < x) +<-asym {.(suc _)} {.(suc _)} (s≤s p) (s≤s q) = <-asym p q + +<-not-equal : ∀ {x y} → x < y → x ≠ y +<-not-equal {zero} (s≤s p) q = absurd (zero≠suc q) +<-not-equal {suc x} (s≤s p) q = <-not-equal p (suc-inj q) + +<-irrefl : ∀ {x y} → x ≡ y → ¬ (x < y) +<-irrefl {suc x} {zero} p q = absurd (suc≠zero p) +<-irrefl {zero} {suc y} p _ = absurd (zero≠suc p) +<-irrefl {suc x} {suc y} p (s≤s q) = <-irrefl (suc-inj p) q + +<-weaken : ∀ {x y} → x < y → x ≤ y +<-weaken {x} {suc y} p = ≤-sucr (≤-peel p) + +≤-strengthen : ∀ {x y} → x ≤ y → (x ≡ y) ⊎ (x < y) +≤-strengthen {zero} {zero} 0≤x = inl refl +≤-strengthen {zero} {suc y} 0≤x = inr (s≤s 0≤x) +≤-strengthen {suc x} {suc y} (s≤s p) with ≤-strengthen p +... | inl eq = inl (ap suc eq) +... | inr le = inr (s≤s le) + +<-from-≤ : ∀ {x y} → x ≠ y → x ≤ y → x < y +<-from-≤ x≠y x≤y with ≤-strengthen x≤y +... | inl x=y = absurd (x≠y x=y) +... | inr x + +```agda + ≤-dec : (x y : Nat) → Dec (x ≤ y) + ≤-dec zero zero = yes 0≤x + ≤-dec zero (suc y) = yes 0≤x + ≤-dec (suc x) zero = no λ { () } + ≤-dec (suc x) (suc y) with ≤-dec x y + ... | yes x≤y = yes (s≤s x≤y) + ... | no ¬x≤y = no (λ { (s≤s x≤y) → ¬x≤y x≤y }) ≤-is-weakly-total : ∀ x y → ¬ (x ≤ y) → y ≤ x ≤-is-weakly-total zero zero _ = 0≤x @@ -87,6 +153,33 @@ Furthermore, `_≤_`{.Agda} is decidable, and weakly total: + + + +```agda +module Data.Nat.Prime where +``` + +# Prime and composite numbers + +A number $n$ is **prime** when it is greater than two (which we split +into being positive and being apart from 1), and, if $d$ is any other +number that divides $n$, then either $d = 1$ or $d = n$. Since we've +assumed $n \neq 1$, these cases are disjoint, so being a prime is a +[[proposition]]. + +```agda +record is-prime (n : Nat) : Type where + field + ⦃ positive ⦄ : Positive n + prime≠1 : n ≠ 1 + primality : ∀ d → d ∣ n → (d ≡ 1) ⊎ (d ≡ n) +``` + +We note that a prime number has no *proper divisors*, in that, if $d +\neq 1$ and $d \neq n$, then $d$ does not divide $n$. + +```agda + ¬proper-divisor : ∀ {d} → d ≠ 1 → d ≠ n → ¬ d ∣ n + ¬proper-divisor ¬d=1 ¬d=n d∣n with primality _ d∣n + ... | inl d=1 = ¬d=1 d=1 + ... | inr d=n = ¬d=n d=n +``` + + + +Conversely, if $n > 2$ is a number that is not prime, it is called +**composite**. Instead of defining `is-composite`{.Agda} to be the +literal negation of `is-prime`{.Agda}, we instead define this notion +positively: To say that a number is composite is to give a prime $p$ and +a number $q \ge 2$ such that $qp = n$ and $p$ is the smallest divisor of +$n$. + +```agda +record is-composite (n : Nat) : Type where + field + {p q} : Nat + + p-prime : is-prime p + q-proper : 1 < q + + factors : q * p ≡ n + least : ∀ p' → 1 < p' → p' ∣ n → p ≤ p' +``` + +Note that the assumption that $p$ is a prime is simply for convenience: +the least proper divisor of *any* number is always a prime. + +```agda +least-divisor→is-prime + : ∀ p n → 1 < p → p ∣ n → (∀ p' → 1 < p' → p' ∣ n → p ≤ p') → is-prime p +least-divisor→is-prime p n gt div least .positive = ≤-trans ≤-ascend gt +least-divisor→is-prime p n gt div least .prime≠1 q = <-irrefl (sym q) gt +least-divisor→is-prime p@(suc p') n gt div least .primality 0 x = absurd (suc≠zero x) +least-divisor→is-prime p@(suc p') n gt div least .primality 1 x = inl refl +least-divisor→is-prime p@(suc p') n gt div least .primality m@(suc (suc k)) x = + inr (≤-antisym (m∣sn→m≤sn x) (least m (s≤s (s≤s 0≤x)) (∣-trans x div))) +``` + + + +## Primality testing + +```agda +is-prime-or-composite : ∀ n → 1 < n → is-prime n ⊎ is-composite n +is-prime-or-composite n@(suc (suc m)) (s≤s p) + with Fin-omniscience {n = n} (λ k → 1 < to-nat k × to-nat k ∣ n) +... | inr prime = inl record { prime≠1 = suc≠zero ∘ suc-inj ; primality = no-divisors→prime } where + no-divisors→prime : ∀ d → d ∣ n → d ≡ 1 ⊎ d ≡ n + no-divisors→prime d div with d ≡? 1 + ... | yes p = inl p + ... | no ¬d=1 with d ≡? n + ... | yes p = inr p + ... | no ¬d=n = + let + ord : d < n + ord = proper-divisor-< ¬d=n div + + coh : d ≡ to-nat (from-ℕ< (d , ord)) + coh = sym (ap fst (to-from-ℕ< (d , ord))) + + no = case d return (λ x → ¬ x ≡ 1 → x ∣ n → 1 < x) of λ where + 0 p1 div → absurd (<-irrefl (sym div) (s≤s 0≤x)) + 1 p1 div → absurd (p1 refl) + (suc (suc n)) p1 div → s≤s (s≤s 0≤x) + in absurd (prime (from-ℕ< (d , ord)) ((subst (1 <_) coh (no ¬d=1 div)) , subst (_∣ n) coh div)) +... | inl (ix , (proper , div) , least) + = inr record { p-prime = prime ; q-proper = proper' ; factors = path ; least = least' } where + open Σ (∣→fibre div) renaming (fst to quot ; snd to path) + + abstract + least' : (p' : Nat) → 1 < p' → p' ∣ n → to-nat ix ≤ p' + least' p' x div with ≤-strengthen (m∣n→m≤n div) + ... | inl same = ≤-trans ≤-ascend (subst (to-nat ix <_) (sym same) (to-ℕ< ix .snd)) + ... | inr less = + let + it : ℕ< n + it = p' , less + + coh : p' ≡ to-nat (from-ℕ< it) + coh = sym (ap fst (to-from-ℕ< it)) + + orig = least (from-ℕ< it) (subst (1 <_) coh x , subst (_∣ n) coh div) + in ≤-trans orig (subst (to-nat (from-ℕ< it) ≤_) (sym coh) ≤-refl) + + prime : is-prime (to-nat ix) + prime = least-divisor→is-prime (to-nat ix) n proper div least' + + proper' : 1 < quot + proper' with quot | path + ... | 0 | q = absurd (<-irrefl q (s≤s 0≤x)) + ... | 1 | q = absurd (<-irrefl (sym (+-zeror (to-nat ix)) ∙ q) (to-ℕ< ix .snd)) + ... | suc (suc n) | p = s≤s (s≤s 0≤x) + +record Factorisation (n : Nat) : Type where + constructor factored + field + primes : List Nat + factors : product primes ≡ n + is-primes : All is-prime primes + + prime-divides : ∀ {x} → x ∈ₗ primes → x ∣ n + prime-divides {x} h = subst (x ∣_) factors (work _ _ h) where + work : ∀ x xs → x ∈ₗ xs → x ∣ product xs + work x (y ∷ xs) (here p) = subst (λ e → x ∣ e * product xs) p (∣-*l {x} {product xs}) + work x (y ∷ xs) (there p) = + let + it = work x xs p + (div , p) = ∣→fibre it + in fibre→∣ (y * div , *-associative y div x ∙ ap (y *_) p) + + find-prime-factor : ∀ {x} → is-prime x → x ∣ n → x ∈ₗ primes + find-prime-factor {num} x d = + work _ primes x (λ _ → all-∈ is-primes) (subst (num ∣_) (sym factors) d) + where + work : ∀ x xs → is-prime x → (∀ x → x ∈ₗ xs → is-prime x) → x ∣ product xs → x ∈ₗ xs + work x [] xp ps xd = absurd (prime≠1 xp (∣-1 xd)) + work x (y ∷ xs) xp ps xd with x ≡? y + ... | yes x=y = here x=y + ... | no ¬p = there $ work x xs xp (λ x w → ps x (there w)) + (|-*-coprime-cancel x y (product xs) + ⦃ auto ⦄ ⦃ ps y (here refl) .positive ⦄ + xd (distinct-primes→coprime xp (ps y (here refl)) ¬p)) + +open is-composite using (factors) +open Factorisation + +factorisation-unique' : ∀ {n} (a b : Factorisation n) → ∀ p → p ∈ₗ a .primes → p ∈ₗ b .primes +factorisation-unique' a b p p∈a = + let + p∣n = prime-divides a p∈a + in find-prime-factor b (all-∈ (a .is-primes) p∈a) p∣n + +factorisation-worker + : ∀ n → (∀ k → k < n → .⦃ _ : Positive k ⦄ → Factorisation k) + → .⦃ _ : Positive n ⦄ → Factorisation n +factorisation-worker 1 ind = factored [] refl [] +factorisation-worker n@(suc (suc m)) ind with is-prime-or-composite n (s≤s (s≤s 0≤x)) +... | inl prime = factored (n ∷ []) (ap (2 +_) (*-oner m)) (prime ∷ []) +... | inr composite@record{p = p; q = q} = + let + instance + _ : Positive q + _ = ≤-trans ≤-ascend (composite .q-proper) + + factored ps path primes = ind q $ + prime-divisor-lt p q n (composite .p-prime) (composite .factors) + in record + { primes = p ∷ ps + ; factors = ap (p *_) path ·· *-commutative p q ·· composite .factors + ; is-primes = composite .p-prime ∷ primes + } + +factorial : Nat → Nat +factorial zero = 1 +factorial (suc n) = suc n * factorial n + +factorial-positive : ∀ n → Positive (factorial n) +factorial-positive zero = s≤s 0≤x +factorial-positive (suc n) = *-preserves-≤ 1 (suc n) 1 (factorial n) (s≤s 0≤x) (factorial-positive n) + +≤-factorial : ∀ n → n ≤ factorial n +≤-factorial zero = 0≤x +≤-factorial (suc n) = subst (_≤ factorial (suc n)) (*-oner (suc n)) (*-preserves-≤ (suc n) (suc n) 1 (factorial n) ≤-refl (factorial-positive n)) + +∣-factorial : ∀ n {m} → m < n → suc m ∣ factorial n +∣-factorial (suc n) {m} m≤n with suc m ≡? suc n +... | yes m=n = subst (_∣ factorial (suc n)) (sym m=n) (∣-*l {suc n} {factorial n}) +... | no m≠n = |-*l-pres {suc m} {suc n} {factorial n} $ + ∣-factorial n {m} (≤-uncap (suc m) n m≠n m≤n) + +factorise : (n : Nat) .⦃ _ : Positive n ⦄ → Factorisation n +factorise = Wf-induction _<_ <-wf _ factorisation-worker + +new-prime : (n : Nat) → Σ[ p ∈ Nat ] n < p × is-prime p +new-prime n with is-prime-or-composite (suc (factorial n)) (s≤s (factorial-positive n)) +new-prime n | inl prime = suc (factorial n) , s≤s (≤-factorial n) , prime +new-prime n | inr c@record{p = p} with holds? (suc n ≤ p) +new-prime n | inr c@record{p = p} | yes n

_ + infixr 8 _<>_ field mempty : A diff --git a/src/Prim/Data/Nat.lagda.md b/src/Prim/Data/Nat.lagda.md index fb050e9c4..cbdf4496b 100644 --- a/src/Prim/Data/Nat.lagda.md +++ b/src/Prim/Data/Nat.lagda.md @@ -31,8 +31,8 @@ are computed more efficiently when bound as `BUILTIN`s. ```agda infix 4 _==_ _<_ -infixl 6 _+_ _-_ -infixl 7 _*_ +infixl 8 _+_ _-_ +infixl 9 _*_ _+_ : Nat → Nat → Nat zero + m = m diff --git a/src/Prim/Data/Sigma.lagda.md b/src/Prim/Data/Sigma.lagda.md index 2b9c3638c..f3ae44856 100644 --- a/src/Prim/Data/Sigma.lagda.md +++ b/src/Prim/Data/Sigma.lagda.md @@ -36,7 +36,7 @@ infixr 4 _,_ Σ-syntax X F = Σ X F syntax Σ-syntax X (λ x → F) = Σ[ x ∈ X ] F -infix 5 Σ-syntax +infix 4 Σ-syntax ``` --> diff --git a/src/Prim/Kan.lagda.md b/src/Prim/Kan.lagda.md index 8010be0cc..fcb20efa3 100644 --- a/src/Prim/Kan.lagda.md +++ b/src/Prim/Kan.lagda.md @@ -50,7 +50,7 @@ postulate {-# BUILTIN PATHP PathP #-} -infix 4 _≡_ +infix 7 _≡_ _≡_ : ∀ {ℓ} {A : Type ℓ} → A → A → Type ℓ _≡_ {A = A} = PathP (λ i → A) @@ -60,6 +60,8 @@ _≡_ {A = A} = PathP (λ i → A) diff --git a/src/1Lab/Type/Pi.lagda.md b/src/1Lab/Type/Pi.lagda.md index 11030a05e..0fc2340c3 100644 --- a/src/1Lab/Type/Pi.lagda.md +++ b/src/1Lab/Type/Pi.lagda.md @@ -1,6 +1,7 @@ @@ -56,6 +52,13 @@ record is-abelian-group (_*_ : G → G → G) : Type (level-of G) where open is-group has-is-group renaming (unit to 1g) public ``` + + -The fundamental example of abelian group is the integers, $\ZZ$, under -addition. A type-theoretic interjection is necessary: the integers live -on the zeroth universe, so to have an $\ell$-sized group of integers, we -must lift it. +The fundamental example of an abelian group is the [[group of integers]]. + +:::{.definition #negation-automorphism} +Given an abelian group $G$, we can define the **negation automorphism** +$G \cong G$ which inverts every element: since the group operation is +commutative, we have $(x \star y)^{-1} = y^{-1} \star x^{-1} = x^{-1} +\star y^{-1}$, so this is a homomorphism. +::: + + ```agda -ℤ-ab : ∀ {ℓ} → Abelian-group ℓ -ℤ-ab = to-ab mk-ℤ where - open make-abelian-group - mk-ℤ : make-abelian-group (Lift _ Int) - mk-ℤ .ab-is-set = hlevel 2 - mk-ℤ .mul (lift x) (lift y) = lift (x +ℤ y) - mk-ℤ .inv (lift x) = lift (negℤ x) - mk-ℤ .1g = lift 0 - mk-ℤ .idl (lift x) = ap lift (+ℤ-zerol x) - mk-ℤ .assoc (lift x) (lift y) (lift z) = ap lift (+ℤ-assoc x y z) - mk-ℤ .invl (lift x) = ap lift (+ℤ-invl x) - mk-ℤ .comm (lift x) (lift y) = ap lift (+ℤ-commutative x y) - -ℤ : ∀ {ℓ} → Group ℓ -ℤ = Abelian→Group ℤ-ab + negation : G Ab.≅ G + negation = total-iso + (_⁻¹ , is-involutive→is-equiv (λ _ → inv-inv)) + (record { pres-⋆ = λ x y → inv-comm ∙ commutes }) ``` diff --git a/src/Algebra/Group/Action.lagda.md b/src/Algebra/Group/Action.lagda.md index 983121d93..62e949ddb 100644 --- a/src/Algebra/Group/Action.lagda.md +++ b/src/Algebra/Group/Action.lagda.md @@ -1,10 +1,13 @@ -# Group actions +# Group actions {defines="group-action"} A useful way to think about [groups] is to think of their elements as encoding "symmetries" of a particular object. For a concrete example, @@ -50,21 +54,17 @@ $\bb{C} \to \bb{C}$ in a way _compatible_ with the group structure: Additive inverses "translate to" inverse maps, addition translates to function composition, and the additive identity is mapped to the identity function. Note that since $\bb{C}$ is a set, this is -equivalently a [group homomorphism] +equivalently a [[group homomorphism]] $$ \bb{R} \to \rm{Sym}(\bb{C}) $$, -where the codomain is the [group of symmetries] of $\bb{C}$. But what if +where the codomain is the [[group of symmetries|symmetric group]] of $\bb{C}$. But what if we want a group $G$ to act on an object $X$ of a more general -[category], rather than an object of $\Sets$? +[[category]], rather than an object of $\Sets$? -[group homomorphism]: Algebra.Group.html#group-homomorphisms -[group of symmetries]: Algebra.Group.html#symmetric-groups -[category]: Cat.Base.html - -## Automorphism groups +## Automorphism groups {defines="automorphism-group"} The answer is that, for an object $X$ of some category $\cC$, the collection of all [isomorphisms] $X \cong X$ forms a group under @@ -83,7 +83,7 @@ module _ {o ℓ} (C : Precategory o ℓ) where mg : make-group (X C.≅ X) mg .make-group.group-is-set = hlevel 2 mg .make-group.unit = C.id-iso - mg .make-group.mul g f = g C.∘Iso f + mg .make-group.mul f g = g C.∘Iso f mg .make-group.inv = C._Iso⁻¹ mg .make-group.assoc x y z = ext (sym (C.assoc _ _ _)) mg .make-group.invl x = ext (x .C.invl) @@ -99,6 +99,14 @@ $G$; A **$G$-action on $X$** is a group homomorphism $G \to Action G X = Groups.Hom G (Aut X) ``` +::: note +Since we've defined $f \star g = g \circ f$ in the automorphism group, +our definition corresponds to *right* actions. If we had defined +$f \star g = f \circ g$ instead, we would get *left* actions. +Of course, the two definitions are formally dual, so it does not +really matter. +::: + # As functors Recall that we defined the [delooping] of a [monoid] $M$ into a category @@ -153,3 +161,50 @@ applying the right helpers for pushing paths inwards, we're left with (is-prop→pathp (λ i → is-group-hom-is-prop) _ _) .is-iso.linv x → Functor-path (λ _ → refl) λ _ → refl ``` + +# Examples of actions + +:::{.definition #trivial-action} +For any group $G$, category $\cC$ and object $X : \cC$, we have the +**trivial action** of $G$ on $X$ which maps every element to the +identity automorphism. It can be defined simply as the [[zero morphism]] +$G \to \rm{Aut}(X)$. +::: + +```agda +module _ {o ℓ} {C : Precategory o ℓ} (G : Group ℓ) where + trivial-action : ∀ X → Action C G X + trivial-action X = Zero.zero→ ∅ᴳ +``` + +:::{.definition #principal-action} +Any group $G$ acts on itself on the right in two ways: first, $G$ acts on its +underlying set by multiplication on the right. This is sometimes called +the **principal action** or the **(right-)regular action**, and is the +basis for [[Cayley's theorem]]. +::: + + + +```agda + principal-action : Action (Sets ℓ) G (G .fst) + principal-action .hom x = equiv→iso ((G._⋆ x) , G.⋆-equivr x) + principal-action .preserves .pres-⋆ x y = ext λ z → G.associative +``` + +$G$ also acts on itself *as a group* by **conjugation**. An automorphism +of $G$ that arises from conjugation with an element of $G$ is called an +**inner automorphism**. + +```agda + conjugation-action : Action (Groups ℓ) G G + conjugation-action .hom x = total-iso + ((λ y → x G.⁻¹ G.⋆ y G.⋆ x) , ∙-is-equiv (G.⋆-equivr x) (G.⋆-equivl (x G.⁻¹))) + (record { pres-⋆ = λ y z → group! G }) + conjugation-action .preserves .pres-⋆ x y = ext λ z → group! G +``` diff --git a/src/Algebra/Group/Cat/Base.lagda.md b/src/Algebra/Group/Cat/Base.lagda.md index a71952259..2943014ab 100644 --- a/src/Algebra/Group/Cat/Base.lagda.md +++ b/src/Algebra/Group/Cat/Base.lagda.md @@ -1,7 +1,10 @@ + ## The underlying set The category of groups admits a `faithful`{.Agda ident=is-faithful} diff --git a/src/Algebra/Group/Cat/FinitelyComplete.lagda.md b/src/Algebra/Group/Cat/FinitelyComplete.lagda.md index a107d50d4..3a89e8ed9 100644 --- a/src/Algebra/Group/Cat/FinitelyComplete.lagda.md +++ b/src/Algebra/Group/Cat/FinitelyComplete.lagda.md @@ -86,7 +86,7 @@ Zero-group-is-zero = record ∅ᴳ .Zero.has-is-zero = Zero-group-is-zero ``` -## Direct products +## Direct products {defines="direct-product-of-groups"} We compute the product of two groups $G \times H$ as the product of their underlying sets, equipped with the operation of "pointwise diff --git a/src/Algebra/Group/Cat/Monadic.lagda.md b/src/Algebra/Group/Cat/Monadic.lagda.md index c5dc14bfe..6c0df60e4 100644 --- a/src/Algebra/Group/Cat/Monadic.lagda.md +++ b/src/Algebra/Group/Cat/Monadic.lagda.md @@ -163,7 +163,7 @@ but the other direction is by induction on "words". ```agda isom : is-iso K.₀ isom .is-iso.inv (A , alg) = A , Algebra-on→group-on alg - isom .is-iso.linv x = ∫-Path Groups-equational + isom .is-iso.linv x = ∫-Path (total-hom (λ x → x) (record { pres-⋆ = λ x y → refl })) id-equiv isom .is-iso.rinv x = Σ-pathp refl (Algebra-on-pathp _ _ go) where diff --git a/src/Algebra/Group/Cayley.lagda.md b/src/Algebra/Group/Cayley.lagda.md index 0022b457a..24db969e3 100644 --- a/src/Algebra/Group/Cayley.lagda.md +++ b/src/Algebra/Group/Cayley.lagda.md @@ -2,6 +2,7 @@ ``` open import 1Lab.Prelude +open import Algebra.Group.Instances.Symmetric open import Algebra.Group.Cat.Base open import Algebra.Group ``` @@ -13,16 +14,14 @@ module Algebra.Group.Cayley {ℓ} (G : Group ℓ) where open Group-on (G .snd) renaming (underlying-set to G-set) ``` -# Cayley's theorem +# Cayley's theorem {defines="cayleys-theorem"} Cayley's theorem says that any group $G$ admits a representation as a -subgroup of a [symmetric group], specifically the symmetric group acting +subgroup of a [[symmetric group]], specifically the symmetric group acting on the underlying set of $G$. -[symmetric group]: Algebra.Group.html#symmetric-groups - First, recall that we get a family of equivalences $G \simeq G$ by multiplication -on the left: +on the left, the [[principal action]] of $G$ on itself: ```agda Cayley : ⌞ G ⌟ → ⌞ G ⌟ ≃ ⌞ G ⌟ diff --git a/src/Algebra/Group/Concrete/Abelian.lagda.md b/src/Algebra/Group/Concrete/Abelian.lagda.md index af05780df..28ded9d8c 100644 --- a/src/Algebra/Group/Concrete/Abelian.lagda.md +++ b/src/Algebra/Group/Concrete/Abelian.lagda.md @@ -3,6 +3,7 @@ open import 1Lab.Path.Reasoning open import 1Lab.Prelude +open import Algebra.Group.Instances.Integers open import Algebra.Group.Cat.Base open import Algebra.Group.Concrete open import Algebra.Group.Ab diff --git a/src/Algebra/Group/Concrete/FinitelyComplete.lagda.md b/src/Algebra/Group/Concrete/FinitelyComplete.lagda.md new file mode 100644 index 000000000..e90691af5 --- /dev/null +++ b/src/Algebra/Group/Concrete/FinitelyComplete.lagda.md @@ -0,0 +1,84 @@ + + +```agda +module Algebra.Group.Concrete.FinitelyComplete {ℓ} where +``` + +# Finite limits of concrete groups + +Since the category of [[concrete groups]] is equivalent to the +[[category of groups]], and the latter is [[finitely complete]], then +so is the former. + +```agda +ConcreteGroups-finitely-complete : Finitely-complete (ConcreteGroups ℓ) +ConcreteGroups-finitely-complete = subst Finitely-complete + (sym (eqv→path ConcreteGroups-is-category Groups-is-category _ π₁F-is-equivalence)) + Groups-finitely-complete +``` + +However, the construction of binary [[products]] of concrete groups +(corresponding to the [[direct product of groups]]) is natural enough to +spell out explicitly: one can simply take the product of the underlying +groupoids! + +```agda +Direct-product-concrete : ConcreteGroup ℓ → ConcreteGroup ℓ → ConcreteGroup ℓ +Direct-product-concrete G H .B = B G ×∙ B H +Direct-product-concrete G H .has-is-connected = is-connected→is-connected∙ $ + ×-is-n-connected 2 + (is-connected∙→is-connected (G .has-is-connected)) + (is-connected∙→is-connected (H .has-is-connected)) +Direct-product-concrete G H .has-is-groupoid = ×-is-hlevel 3 + (G .has-is-groupoid) + (H .has-is-groupoid) + +ConcreteGroups-products + : (X Y : ConcreteGroup ℓ) → Product (ConcreteGroups ℓ) X Y +ConcreteGroups-products X Y = prod where + prod : Product (ConcreteGroups ℓ) X Y + prod .apex = Direct-product-concrete X Y + prod .π₁ = fst∙ + prod .π₂ = snd∙ + prod .has-is-product .⟨_,_⟩ f g .fst x = f # x , g # x + prod .has-is-product .⟨_,_⟩ f g .snd = f .snd ,ₚ g .snd + prod .has-is-product .π₁∘⟨⟩ = funext∙ (λ _ → refl) (∙-idr _) + prod .has-is-product .π₂∘⟨⟩ = funext∙ (λ _ → refl) (∙-idr _) + prod .has-is-product .unique {Q} {f} {g} {u} p1 p2 = + funext∙ (λ x → p1 #ₚ x ,ₚ p2 #ₚ x) (fix ◁ square) + where + square : Square + (p1 #ₚ pt Q ,ₚ p2 #ₚ pt Q) ((fst∙ ∘∙ u) .snd ,ₚ (snd∙ ∘∙ u) .snd) + (f .snd ,ₚ g .snd) refl + square i = p1 i .snd ,ₚ p2 i .snd + + fix : u .snd ≡ ((fst∙ ∘∙ u) .snd ,ₚ (snd∙ ∘∙ u) .snd) + fix i = ∙-idr (ap fst (u .snd)) (~ i) ,ₚ ∙-idr (ap snd (u .snd)) (~ i) +``` + +Similarly, the [[terminal]] concrete group is just the unit type. + +```agda +Terminal-concrete : ConcreteGroup ℓ +Terminal-concrete .B = Lift _ ⊤ , _ +Terminal-concrete .has-is-connected = is-connected→is-connected∙ (n-connected 2) +Terminal-concrete .has-is-groupoid = hlevel 3 +``` diff --git a/src/Algebra/Group/Concrete/Semidirect.lagda.md b/src/Algebra/Group/Concrete/Semidirect.lagda.md new file mode 100644 index 000000000..c0d3e7787 --- /dev/null +++ b/src/Algebra/Group/Concrete/Semidirect.lagda.md @@ -0,0 +1,207 @@ + + +```agda +module Algebra.Group.Concrete.Semidirect {ℓ} where +``` + +# Semidirect products of concrete groups {defines="semidirect-product-of-concrete-groups"} + +The construction of the [[semidirect product]] for [[concrete groups]] +is quite natural: given a concrete group $G$ and an action $\varphi : +\B G \to \rm{ConcreteGroups}$ of $G$ on $H = \varphi(\point{G})$, we +simply define $\B(G \ltimes H)$ as the $\Sigma$-type + +$$ +\sum_{g : \B G} \B \varphi(g) +$$ + +pointed at $(\point{G}, \point{H})$. + + + +```agda + H : ConcreteGroup ℓ + H = φ (pt G) + + Semidirect-concrete : ConcreteGroup ℓ + Semidirect-concrete .B .fst = Σ ⌞ B G ⌟ λ g → ⌞ B (φ g) ⌟ + Semidirect-concrete .B .snd = pt G , pt H +``` + +That this is a connected groupoid follows from standard results. + +```agda + Semidirect-concrete .has-is-connected = is-connected→is-connected∙ $ + Σ-is-n-connected 2 + (is-connected∙→is-connected (G .has-is-connected)) + (λ a → is-connected∙→is-connected (φ a .has-is-connected)) + Semidirect-concrete .has-is-groupoid = Σ-is-hlevel 3 + (G .has-is-groupoid) + (λ a → φ a .has-is-groupoid) +``` + +In order to relate this with the definition of [[semidirect products]] +for abstract groups, we show that the [[fundamental group]] +$\pi_1\B(G \ltimes H)$ is the semidirect product $\pi_1\B G \ltimes +\pi_1\B H$. + +To that end, we start by defining the action of $\pi_1\B G$ on +$\pi_1\B H$ induced by $\varphi$. +Given a loop $p : \point{G} \equiv \point{G}$, we can apply the function +$x \mapsto \pi_1 \B \varphi(x)$ to it to get a loop $\pi_1 \B H \equiv +\pi_1 \B H$, which we then turn into an automorphism. + + + +```agda + π₁-φ : Action (Groups ℓ) (π₁B G) (π₁B H) + π₁-φ .hom p = path→iso (ap (π₁B ⊙ φ) p) + π₁-φ .preserves .pres-⋆ p q = + path→iso (ap (π₁B ⊙ φ) (p ∙ q)) ≡⟨ ap path→iso (ap-∙ (π₁B ⊙ φ) p q) ⟩ + path→iso (ap (π₁B ⊙ φ) p ∙ ap (π₁B ⊙ φ) q) ≡⟨ path→iso-∙ (Groups ℓ) _ _ ⟩ + path→iso (ap (π₁B ⊙ φ) p) ∙Iso path→iso (ap (π₁B ⊙ φ) q) ∎ + + π₁BG⋉π₁BH : Group ℓ + π₁BG⋉π₁BH = Semidirect-product (π₁B G) (π₁B H) π₁-φ + + private module ⋉ = Group-on (π₁BG⋉π₁BH .snd) +``` + +::: source +The following proof is adapted from Symmetry [-@Symmetry, §4.14]. +::: + +The first step is to construct an equivalence between the underlying +sets of $\pi_1\B(G \ltimes H)$ and $\pi_1\B G \ltimes \pi_1\B H$. +Since $\B(G \ltimes H)$ is a $\Sigma$-type, its loops consist of pairs +$(p : \point{G} \is \point{G}, q : \point{H} \is_p \point{H})$, where +the second component is a [[dependent path]] over $p$ in the total space +of $\varphi$. What we actually want is a pair of non-dependent loops, +so we need to "realign" the second component by transporting one of +its endpoints along $p$; we do this by generalising on the endpoints +and using [[path induction]]. + +Note that we have a type $A = \B G$, a type family $B(a) = \B\varphi(a)$, +and a function $f : (a : A) \to B(a)$ given by $a \mapsto \point{\varphi(a)}$. +The following constructions work in this general setting. + + + +```agda + align + : ∀ {a a' : A} (p : a ≡ a') + → PathP (λ i → B (p i)) (f a) (f a') ≃ (f a' ≡ f a') + align {a} = + J (λ a' p → PathP (λ i → B (p i)) (f a) (f a') ≃ (f a' ≡ f a')) id≃ + + align-refl + : ∀ {a : A} (q : f a ≡ f a) + → align (λ _ → a) # q ≡ q + align-refl {a} = unext $ + J-refl (λ a' p → PathP (λ i → B (p i)) (f a) (f a') ≃ (f a' ≡ f a')) id≃ + + Σ-align + : ∀ {a a' : A} + → ((a , f a) ≡ (a' , f a')) ≃ ((a ≡ a') × (f a' ≡ f a')) + Σ-align = Iso→Equiv Σ-pathp-iso e⁻¹ ∙e Σ-ap-snd align +``` + +It remains to show that this equivalence is a [[group homomorphism]] +--- or rather a group*oid* homomorphism, since we've generalised the +endpoints. +First, note that we can describe the action of $\varphi$ on paths using +a kind of dependent [[conjugation]] operation. + +```agda + conjP : ∀ {a a' : A} → (p : a ≡ a') → f a ≡ f a → f a' ≡ f a' + conjP {a} p = transport (λ i → f a ≡ f a → f (p i) ≡ f (p i)) λ q → q + + conjP-refl : ∀ {a : A} → (q : f a ≡ f a) → conjP refl q ≡ q + conjP-refl {a} = unext $ transport-refl {A = f a ≡ f a → _} λ q → q +``` + +Thus we have a composition operation on pairs of paths that generalises +the multiplication in $\pi_1\B G \ltimes \pi_1\B H$. + +```agda + _∙⋉_ + : ∀ {a₀ a₁ a₂ : A} + → (a₀ ≡ a₁) × (f a₁ ≡ f a₁) + → (a₁ ≡ a₂) × (f a₂ ≡ f a₂) + → (a₀ ≡ a₂) × (f a₂ ≡ f a₂) + (p , q) ∙⋉ (p' , q') = p ∙ p' , conjP p' q ∙ q' +``` + +

+ +We can then show that `Σ-align`{.Agda} preserves the group(oid) +operations by double path induction. + +```agda + Σ-align-∙ + : ∀ {a₀ a₁ a₂ : A} + → (p : (a₀ , f a₀) ≡ (a₁ , f a₁)) + → (p' : (a₁ , f a₁) ≡ (a₂ , f a₂)) + → Σ-align # (p ∙ p') + ≡ Σ-align # p ∙⋉ Σ-align # p' +``` + + +```agda + Σ-align-∙ {a₀} {a₁} {a₂} p p' = J₂ + (λ a₀ a₂ p⁻¹ p' → + ∀ (q : PathP (λ i → B (p⁻¹ (~ i))) (f a₀) (f a₁)) + → (q' : PathP (λ i → B (p' i)) (f a₁) (f a₂)) + → Σ-align # ((sym p⁻¹ ,ₚ q) ∙ (p' ,ₚ q')) + ≡ Σ-align # (sym p⁻¹ ,ₚ q) ∙⋉ Σ-align # (p' ,ₚ q')) + (λ q q' → + Σ-align # ((refl ,ₚ q) ∙ (refl ,ₚ q')) ≡˘⟨ ap# Σ-align (ap-∙ (a₁ ,_) q q') ⟩ + Σ-align # (refl ,ₚ q ∙ q') ≡⟨⟩ + refl , align refl # (q ∙ q') ≡⟨ refl ,ₚ align-refl (q ∙ q') ⟩ + refl , q ∙ q' ≡˘⟨ refl ,ₚ ap₂ _∙_ (align-refl q) (align-refl q') ⟩ + refl , align refl # q ∙ align refl # q' ≡˘⟨ ∙-idl _ ,ₚ ap (_∙ align refl # q') (conjP-refl (align refl # q)) ⟩ + refl ∙ refl , conjP refl (align refl # q) ∙ align refl # q' ∎) + (ap fst (sym p)) (ap fst p') (ap snd p) (ap snd p') +``` +
+ +This concludes the proof that the constructions of the semidirect +product of abstract and concrete groups agree over the equivalence +between abstract and concrete groups. + +```agda + Semidirect-concrete-abstract : π₁B Semidirect-concrete ≡ π₁BG⋉π₁BH + Semidirect-concrete-abstract = ∫-Path + (total-hom (Σ-align (pt ⊙ φ) .fst) + (record { pres-⋆ = Σ-align-∙ (pt ⊙ φ) })) + (Σ-align (pt ⊙ φ) .snd) +``` diff --git a/src/Algebra/Group/Free/Product.lagda.md b/src/Algebra/Group/Free/Product.lagda.md new file mode 100644 index 000000000..f31e942f3 --- /dev/null +++ b/src/Algebra/Group/Free/Product.lagda.md @@ -0,0 +1,188 @@ + + +```agda +module Algebra.Group.Free.Product {ℓ} where +``` + +# Free products of groups {defines="free-product amalgamated-free-product"} + +Given two [[groups]] $A$ and $B$, their **free product** is the [[coproduct]] +$A + B$ in the [[category of groups]] (not to be confused with the +[[*direct* product|direct product of groups]] $A \times B$, which is the +categorical [[product]]). + +As the name suggests, we can describe the free product $A + B$ as the +free group equipped with inclusions $A \hookrightarrow A + B$ and +$B \hookrightarrow A + B$. In fact, we define a more general construction: +the **amalgamated free product** of a span of groups +$A \overset{f}{\ot} C \overset{g}{\to} B$, which realises the +[[pushout]] of this span. + + + +We start by freely adding the group operations and enforcing the group +axioms, just like for [[free groups|free group construction]]... + +```agda + data Amalgamated : Type ℓ where + _◆_ : Amalgamated → Amalgamated → Amalgamated + inv : Amalgamated → Amalgamated + nil : Amalgamated + + f-assoc : ∀ x y z → x ◆ (y ◆ z) ≡ (x ◆ y) ◆ z + f-invl : ∀ x → inv x ◆ x ≡ nil + f-idl : ∀ x → nil ◆ x ≡ x +``` + +...we then throw in the inclusions of $A$ and $B$, which are required to +be [[group homomorphisms]] and to make the pushout square commute... + +```agda + inl : ⌞ A ⌟ → Amalgamated + inl-hom : ∀ a a' → inl (a A.⋆ a') ≡ inl a ◆ inl a' + + inr : ⌞ B ⌟ → Amalgamated + inr-hom : ∀ b b' → inr (b B.⋆ b') ≡ inr b ◆ inr b' + + glue : (c : ⌞ C ⌟) → inl (f # c) ≡ inr (g # c) +``` + +...finally, we truncate the resulting type to a set. + +```agda + squash : is-set Amalgamated +``` + + + +
+ +As expected, this data perfectly assembles into a group $A +_C B$ +together with homomorphisms $A \to A +_C B$ and $B \to A +_C B$ +fitting into a square with $f$ and $g$. + +~~~{.quiver} +\[\begin{tikzcd} + & C \\ + A && B \\ + & {A +_C B} + \arrow["f"', from=1-2, to=2-1] + \arrow["g", from=1-2, to=2-3] + \arrow["{\mathrm{inl}}"', from=2-1, to=3-2] + \arrow["{\mathrm{inr}}", from=2-3, to=3-2] +\end{tikzcd}\] +~~~ + +```agda + Amalgamated-free-product : Group ℓ + inlᴳ : Groups.Hom A Amalgamated-free-product + inrᴳ : Groups.Hom B Amalgamated-free-product + glueᴳ : inlᴳ Groups.∘ f ≡ inrᴳ Groups.∘ g +``` + + +```agda + Amalgamated-free-product = to-group fp where + fp : make-group Amalgamated + fp .make-group.group-is-set = squash + fp .make-group.unit = nil + fp .make-group.mul = _◆_ + fp .make-group.inv = inv + fp .make-group.assoc = f-assoc + fp .make-group.invl = f-invl + fp .make-group.idl = f-idl + + inlᴳ .hom = inl + inlᴳ .preserves .pres-⋆ = inl-hom + + inrᴳ .hom = inr + inrᴳ .preserves .pres-⋆ = inr-hom + + glueᴳ = ext glue +``` +
+ +The universal property of the pushout is easy to verify. + +```agda + Groups-pushout : is-pushout (Groups ℓ) f inlᴳ g inrᴳ + Groups-pushout .square = glueᴳ + Groups-pushout .universal {Q} {p} {q} comm .hom = go where + module Q = Group-on (Q .snd) + go : Amalgamated → ⌞ Q ⌟ + go (x ◆ y) = go x Q.⋆ go y + go (inv x) = go x Q.⁻¹ + go nil = Q.unit + go (f-assoc x y z i) = Q.associative {go x} {go y} {go z} i + go (f-invl x i) = Q.inversel {go x} i + go (f-idl x i) = Q.idl {go x} i + go (inl a) = p # a + go (inl-hom a a' i) = p .preserves .pres-⋆ a a' i + go (inr b) = q # b + go (inr-hom b b' i) = q .preserves .pres-⋆ b b' i + go (glue c i) = unext comm c i + go (squash x y α β i j) = + Q.has-is-set (go x) (go y) (λ i → go (α i)) (λ i → go (β i)) i j + Groups-pushout .universal comm .preserves .pres-⋆ _ _ = refl + Groups-pushout .universal∘i₁ = trivial! + Groups-pushout .universal∘i₂ = trivial! + Groups-pushout .unique {Q = Q} {colim' = u} comm₁ comm₂ = ext $ + Amalgamated-elim-prop (λ _ → hlevel 1) + (λ x p y q → u .preserves .pres-⋆ x y ∙ ap₂ Q._⋆_ p q) + (λ x p → pres-inv (u .preserves) ∙ ap Q._⁻¹ p) + (pres-id (u .preserves)) + (unext comm₁) (unext comm₂) + where module Q = Group-on (Q .snd) + + Groups-Pushout : Pushout (Groups ℓ) f g + Groups-Pushout .coapex = Amalgamated-free-product + Groups-Pushout .i₁ = inlᴳ + Groups-Pushout .i₂ = inrᴳ + Groups-Pushout .has-is-po = Groups-pushout +``` + +Since the [[zero group]] is also an [[initial object]], this shows that the +category of groups is [[finitely cocomplete]]. + +```agda +Groups-finitely-cocomplete : Finitely-cocomplete (Groups ℓ) +Groups-finitely-cocomplete = with-pushouts (Groups _) + (Zero.initial ∅ᴳ) Groups-Pushout +``` + +We thus get the free product of $A$ and $B$ by abstract nonsense, as the +pushout of the span $A \overset{¡}{\ot} 0 \overset{¡}{\to} B$. + +```agda +Free-product : Group ℓ → Group ℓ → Group ℓ +Free-product A B = Groups-finitely-cocomplete .coproducts A B .coapex +``` diff --git a/src/Algebra/Group/Homotopy.lagda.md b/src/Algebra/Group/Homotopy.lagda.md index d461ea914..1266e7d06 100644 --- a/src/Algebra/Group/Homotopy.lagda.md +++ b/src/Algebra/Group/Homotopy.lagda.md @@ -162,7 +162,7 @@ module π₁Groupoid {ℓ} ((T , t) : Type∙ ℓ) (grpd : is-groupoid T) where π₁ = to-group mk π₁≡π₀₊₁ : π₁ ≡ πₙ₊₁ 0 (T , t) - π₁≡π₀₊₁ = ∫-Path Groups-equational + π₁≡π₀₊₁ = ∫-Path (total-hom inc (record { pres-⋆ = λ _ _ → refl })) (∥-∥₀-idempotent (grpd _ _)) ``` diff --git a/src/Algebra/Group/Homotopy/BAut.lagda.md b/src/Algebra/Group/Homotopy/BAut.lagda.md index 5e57f6be3..ee0ed874f 100644 --- a/src/Algebra/Group/Homotopy/BAut.lagda.md +++ b/src/Algebra/Group/Homotopy/BAut.lagda.md @@ -10,17 +10,14 @@ module Algebra.Group.Homotopy.BAut where # Deloopings of automorphism groups -Recall that any set $X$ generates a group [$\rm{Sym}(X)$][symg], given +Recall that any set $X$ generates a group [[$\rm{Sym}(X)$|symmetric group]] given by the automorphisms $X \simeq X$. We also have a generic construction of [[deloopings]]: special spaces $K(G,1)$ (for a group $G$), where the -[fundamental group] $\pi_1(K(G,1))$ recovers $G$. For the specific case +[[fundamental group]] $\pi_1(K(G,1))$ recovers $G$. For the specific case of deloping automorphism groups, we can give an alternative construction: The type of small types [[merely]] equivalent to $X$ has a fundamental group of $\rm{Sym}(X)$. -[symg]: Algebra.Group.html#symmetric-groups -[fundamental group]: Algebra.Group.Homotopy.html#homotopy-groups - ```agda module _ {ℓ} (T : Type ℓ) where BAut : Type (lsuc ℓ) diff --git a/src/Algebra/Group/Instances/Cyclic.lagda.md b/src/Algebra/Group/Instances/Cyclic.lagda.md new file mode 100644 index 000000000..e45f16113 --- /dev/null +++ b/src/Algebra/Group/Instances/Cyclic.lagda.md @@ -0,0 +1,235 @@ + + +```agda +module Algebra.Group.Instances.Cyclic where +``` + + + +# Cyclic groups {defines="cyclic-group"} + +We say that a [[group]] is **cyclic** if it is generated by a single +element $g$: that is, every element of the group can be written as some +integer power of $g$^[Some authors (namely, Bourbaki) call such groups +*monogeneous*, and reserve "cyclic" for the *finite* cyclic groups.]. + +```agda +_is-generated-by_ : ∀ {ℓ} (G : Group ℓ) (g : ⌞ G ⌟) → Type ℓ +G is-generated-by g = ∀ x → ∃[ n ∈ Int ] x ≡ pow G g n + +is-cyclic : ∀ {ℓ} (G : Group ℓ) → Type ℓ +is-cyclic G = ∃[ g ∈ G ] G is-generated-by g +``` + +::: note +Generators are not unique, so we must use $\exists$ to get a [[proposition]]. +In fact, if $g$ generates $G$, then so does $g^{-1}$, hence a non-trivial +cyclic group has at least two generators! + +```agda +module _ {ℓ} (G : Group ℓ) (open Group-on (G .snd)) where + inverse-generates + : ∀ g → G is-generated-by g → G is-generated-by g ⁻¹ + inverse-generates g gen x = gen x <&> λ (n , p) → negℤ n , ( + x ≡⟨ p ⟩ + pow G g n ≡˘⟨ inv-inv ⟩ + pow G g n ⁻¹ ⁻¹ ≡˘⟨ ap _⁻¹ (pres-inv (pow-hom G g .preserves) {lift n}) ⟩ + pow G g (negℤ n) ⁻¹ ≡˘⟨ pow-inverse G g (negℤ n) ⟩ + pow G (g ⁻¹) (negℤ n) ∎) +``` +::: + +The obvious example is the **infinite cyclic group** $\ZZ$ of the (additive) +integers generated by $1$, but there is also a **finite cyclic group** of order +$n$ for every $n \geq 1$. As it turns out, we can construct *all* the cyclic +groups using the same machine: the quotient group $\ZZ/n\ZZ$, where $n\ZZ$ +is the [[normal subgroup]] of multiples of $n$, is the cyclic group of order +$n$ when $n \geq 1$, and is the infinite cyclic group when $n = 0$. + +```agda +infix 30 _·ℤ +_·ℤ : ∀ (n : Nat) → normal-subgroup ℤ λ i → el (n ∣ℤ i) (∣ℤ-is-prop n i) +(n ·ℤ) .has-rep .has-unit = ∣ℤ-zero +(n ·ℤ) .has-rep .has-⋆ = ∣ℤ-+ +(n ·ℤ) .has-rep .has-inv = ∣ℤ-negℤ +(n ·ℤ) .has-conjugate {x} {y} = subst (n ∣ℤ_) x≡y+x-y + where + x≡y+x-y : x ≡ y +ℤ (x -ℤ y) + x≡y+x-y = + x ≡⟨ ℤ.insertl {h = y} (ℤ.inverser {x = y}) ⟩ + y +ℤ (negℤ y +ℤ x) ≡⟨ ap (y +ℤ_) (+ℤ-commutative (negℤ y) x) ⟩ + y +ℤ (x -ℤ y) ∎ + +infix 25 ℤ/_ +ℤ/_ : Nat → Group lzero +ℤ/ n = ℤ /ᴳ n ·ℤ +``` + +$\ZZ/n\ZZ$ is an [[abelian group]], since the commutativity of +addition descends to the quotient. + +```agda +ℤ/n-commutative : ∀ n → is-commutative-group (ℤ/ n) +ℤ/n-commutative n = elim! λ x y → ap inc (+ℤ-commutative x y) + +infix 25 ℤ-ab/_ +ℤ-ab/_ : Nat → Abelian-group lzero +ℤ-ab/_ n = from-commutative-group (ℤ/ n) (ℤ/n-commutative n) +``` + +Furthermore, $\ZZ/n\ZZ$ is, by construction, cyclic, and is generated +by $1$. + + + +```agda + ℤ/n-cyclic : is-cyclic (ℤ/ n) + ℤ/n-cyclic = inc (1 , gen) where + gen : ℤ/ n is-generated-by 1 + gen = elim! λ x → inc (x , Integers.induction Int-integers + {P = λ x → inc x ≡ pow (ℤ/ n) 1 x} + refl + (λ x → + inc x ≡ pow (ℤ/ n) 1 x ≃⟨ _ , equiv→cancellable (⋆-equivr 1) ⟩ + inc x ⋆ 1 ≡ pow (ℤ/ n) 1 x ⋆ 1 ≃⟨ ∙-pre-equiv (sym (ap inc (+ℤ-oner x))) ⟩ + inc (sucℤ x) ≡ pow (ℤ/ n) 1 x ⋆ 1 ≃⟨ ∙-post-equiv (sym (pow-sucr (ℤ/ n) 1 x)) ⟩ + inc (sucℤ x) ≡ pow (ℤ/ n) 1 (sucℤ x) ≃∎) + x) +``` + +We prove the following mapping-out property for $\ZZ/n\ZZ$: a [[group +homomorphism]] $\ZZ/n\ZZ \to G$ is uniquely determined by an element +$x : G$ of *order* $n$, i.e. such that $x^n = 1$. Note that this +condition is trivially satisfied if $n = 0$, so we don't need a +special case. + +This follows from the fact that a group homomorphism $\ZZ \to G$ out of +the [[group of integers]] is uniquely determined by an element of $G$, +and some basic algebra. + +```agda +module _ + (n : Nat) + {ℓ} {G : Group ℓ} (open Group-on (G .snd)) + (x : ⌞ G ⌟) (open pow G x renaming (pow to infixr 30 x^_; pow-hom to x^-)) + (wraps : x^ pos n ≡ unit) + where + + ℤ/-out : Groups.Hom (LiftGroup ℓ (ℤ/ n)) G + ℤ/-out .hom (lift i) = Coeq-rec (apply x^- ⊙ lift) (λ (a , b , n∣a-b) → zero-diff $ + let k , k*n≡a-b = ∣ℤ→fibre n∣a-b in + x^ a — x^ b ≡˘⟨ pres-diff (x^- .preserves) {lift a} {lift b} ⟩ + x^ (a -ℤ b) ≡˘⟨ ap x^_ k*n≡a-b ⟩ + x^ (k *ℤ pos n) ≡⟨ ap x^_ (*ℤ-commutative k (pos n)) ⟩ + x^ (pos n *ℤ k) ≡⟨ pow-* G x (pos n) k ⟩ + pow G ⌜ x^ pos n ⌝ k ≡⟨ ap! wraps ⟩ + pow G unit k ≡⟨ pow-unit G k ⟩ + unit ∎) + i + ℤ/-out .preserves .pres-⋆ = elim! λ x y → + x^- .preserves .pres-⋆ (lift x) (lift y) +``` + +We can check that $\ZZ/0\ZZ \is \ZZ$: + +```agda +ℤ-ab/0≡ℤ-ab : ℤ-ab/ 0 ≡ ℤ-ab +ℤ-ab/0≡ℤ-ab = ∫-Path + (total-hom + (Coeq-rec (λ z → z) (λ (_ , _ , p) → ℤ.zero-diff p)) + (record { pres-⋆ = elim! λ _ _ → refl })) + (is-iso→is-equiv (iso inc (λ _ → refl) (elim! λ _ → refl))) + +ℤ/0≡ℤ : ℤ/ 0 ≡ ℤ +ℤ/0≡ℤ = Grp→Ab→Grp (ℤ/ 0) (ℤ/n-commutative 0) ∙ ap Abelian→Group ℤ-ab/0≡ℤ-ab +``` + +For positive $n$, we show that $\ZZ/n\ZZ$ is finite; more precisely, +we construct an equivalence $\ZZ/n\ZZ \simeq \NN^{ + +```agda +module Algebra.Group.Instances.Dihedral where +``` + +# Dihedral groups {defines="dihedral-group generalised-dihedral-group"} + +The [[group]] of symmetries of a regular $n$-gon, comprising rotations +around the center and reflections, is called the **dihedral group** +of order $2n$: there are $n$ rotations and $n$ reflections, where +each reflection can be obtained by composing a rotation with a chosen +reflection. + +::: terminology +There are two notational conventions for the dihedral group associated +with the regular $n$-gon: the *geometric* convention is to write $D_n$, +while the *algebraic* convention is to write $D_{2n}$, alluding to the +order of the group. We will use the geometric convention here. +::: + +Given any [[abelian group]] $G$, we define the **generalised dihedral +group** $\mathrm{Dih}(G)$ as the [[semidirect product]]^[See +[[there|semidirect product]] for some intuition.] group $\ZZ/2\ZZ \ltimes +G$, where the [[cyclic group]] $\ZZ/2\ZZ$ [[acts|group action]] on $G$ +by `negation`{.Agda}. + +```agda +Dih : ∀ {ℓ} → Abelian-group ℓ → Group ℓ +Dih G = Semidirect-product (LiftGroup _ (ℤ/ 2)) (Abelian→Group G) $ + ℤ/-out 2 (F-map-iso Ab↪Grp (negation G)) (ext λ _ → inv-inv) + where open Abelian-group-on (G .snd) +``` + +We can then define the dihedral group $D_n$ as the generalised dihedral +group of $\ZZ/n\ZZ$, the group of *rotations* of the regular $n$-gon. + +```agda +D : Nat → Group lzero +D n = Dih (ℤ-ab/ n) +``` + +We also define the **infinite dihedral group** $D_\infty$ as the +generalised dihedral group of the [[group of integers]], +$\mathrm{Dih}(\ZZ)$. + +```agda +D∞ : Group lzero +D∞ = Dih ℤ-ab +``` + +Geometrically, $D_\infty$ is the symmetry group of the following +infinitely repeating frieze pattern, also known as $p1m1$: there is +a horizontal translation for every integer, as well as a reflection +across the vertical axis that flips all the translations. + +~~~{.quiver} +\begin{tikzpicture} + \node (-3) at (-3, 0) {}; + \node (-2) at (-2, 0) {}; + \node (-1) at (-1, 0) {}; + \node (0) at (0, 0) {}; + \node (1) at (1, 0) {}; + \node (2) at (2, 0) {}; + \node (3) at (3, 0) {}; + \draw (-3.center) to (-2.center); + \draw (-2.center) to (-1.center); + \draw (-1.center) to (0.center); + \draw (0.center) to (1.center); + \draw (1.center) to (2.center); + \draw (2.center) to (3.center); + + \node[left,at=(-3)] {$\cdots$}; + \node[right,at=(3)] {$\cdots$}; + + \node (-2') at (-2, 1) {}; + \draw (-2.center) to (-2'.center); + \node (-1') at (-1, 1) {}; + \draw (-1.center) to (-1'.center); + \node (0') at (0, 1) {}; + \draw (0.center) to (0'.center); + \node (1') at (1, 1) {}; + \draw (1.center) to (1'.center); + \node (2') at (2, 1) {}; + \draw (2.center) to (2'.center); +\end{tikzpicture} +~~~ + +Note that our construction of $D_n$ makes $D_0$ equal to $D_\infty$, +since $\ZZ/0\ZZ$ is $\ZZ$. + +```agda +D₀≡D∞ : D 0 ≡ D∞ +D₀≡D∞ = ap Dih ℤ-ab/0≡ℤ-ab +``` diff --git a/src/Algebra/Group/Instances/Integers.lagda.md b/src/Algebra/Group/Instances/Integers.lagda.md new file mode 100644 index 000000000..50718c63e --- /dev/null +++ b/src/Algebra/Group/Instances/Integers.lagda.md @@ -0,0 +1,248 @@ + + +```agda +module Algebra.Group.Instances.Integers where +``` + + + +# The group of integers {defines="integer-group group-of-integers"} + +The fundamental example of an [[abelian group]] is the group of +[[**integers**]], $\ZZ$, under addition. + +```agda +ℤ-ab : Abelian-group lzero +ℤ-ab = to-ab mk-ℤ where + open make-abelian-group + mk-ℤ : make-abelian-group Int + mk-ℤ .ab-is-set = hlevel 2 + mk-ℤ .mul x y = x +ℤ y + mk-ℤ .inv x = negℤ x + mk-ℤ .1g = 0 + mk-ℤ .idl x = +ℤ-zerol x + mk-ℤ .assoc x y z = +ℤ-assoc x y z + mk-ℤ .invl x = +ℤ-invl x + mk-ℤ .comm x y = +ℤ-commutative x y + +ℤ : Group lzero +ℤ = Abelian→Group ℤ-ab +``` + +We show that $\ZZ$ is the [[free group]] on one generator, i.e. the +[[free object]] on $\top$ relative to the forgetful functor from groups +to sets, `Grp↪Sets`{.Agda}. + +We start by defining the [[group homomorphism]] $\ZZ \to G$ that sends +$n$ to $x^n$ (or, in additive notation, $n \cdot x$), where $G$ is a group +and $x$ is an element of $G$, using the [[universal property of the integers]]. + + + +```agda + module pow (x : ⌞ G ⌟) where + pow : Int → ⌞ G ⌟ + pow = ℤ.map-out unit ((_⋆ x) , ⋆-equivr x) + + pow-sucr : ∀ a → pow (sucℤ a) ≡ pow a ⋆ x + pow-sucr = ℤ.map-out-rotate _ _ + + pow-+ : ∀ a b → pow (a +ℤ b) ≡ pow a ⋆ pow b + pow-+ a = ℤ.induction + (ap pow (+ℤ-zeror a) ∙ sym idr) + λ b → + pow (a +ℤ b) ≡ pow a ⋆ pow b ≃⟨ ap (_⋆ x) , equiv→cancellable (⋆-equivr x) ⟩ + pow (a +ℤ b) ⋆ x ≡ (pow a ⋆ pow b) ⋆ x ≃⟨ ∙-post-equiv (sym associative) ⟩ + pow (a +ℤ b) ⋆ x ≡ pow a ⋆ pow b ⋆ x ≃⟨ ∙-post-equiv (ap (pow a ⋆_) (sym (pow-sucr b))) ⟩ + pow (a +ℤ b) ⋆ x ≡ pow a ⋆ pow (sucℤ b) ≃⟨ ∙-pre-equiv (pow-sucr (a +ℤ b)) ⟩ + pow (sucℤ (a +ℤ b)) ≡ pow a ⋆ pow (sucℤ b) ≃⟨ ∙-pre-equiv (ap pow (+ℤ-sucr a b)) ⟩ + pow (a +ℤ sucℤ b) ≡ pow a ⋆ pow (sucℤ b) ≃∎ +``` + +A type-theoretic interjection is necessary: the integers live on the +zeroth universe, so to have an $\ell$-sized group of integers, we +must lift it. + +```agda + pow-hom : Groups.Hom (LiftGroup ℓ ℤ) G + pow-hom .hom (lift i) = pow i + pow-hom .preserves .pres-⋆ (lift a) (lift b) = pow-+ a b +``` + +This is the unique group homomorphism $\ZZ \to G$ that sends $1$ to $x$. + +```agda + pow-unique : (g : Groups.Hom (LiftGroup ℓ ℤ) G) → g # 1 ≡ x → g ≡ pow-hom + pow-unique g g1≡x = ext $ ℤ.map-out-unique (λ i → g # lift i) + (pres-id (g .preserves)) + λ y → + g # lift ⌜ sucℤ y ⌝ ≡⟨ ap! (sym (+ℤ-oner y)) ⟩ + g # lift (y +ℤ 1) ≡⟨ g .preserves .pres-⋆ (lift y) 1 ⟩ + g # lift y ⋆ g # 1 ≡⟨ ap (g # lift y ⋆_) g1≡x ⟩ + g # lift y ⋆ x ∎ + + open pow public +``` + +
+ +We prove some other useful basic properties of exponentiation. +^[Pedantically, these properties say that `pow`{.Agda} is the unique +*near-ring* homomorphism from $\ZZ$, the [[initial near-ring|initial +ring]], to the near-ring of (pointed) endofunctions of $G$, a generalisation +of [[endomorphism rings]] to non-abelian groups.] + +```agda + pow-sucl : ∀ x a → pow x (sucℤ a) ≡ x ⋆ pow x a + pow-0 : ∀ x → pow x 0 ≡ unit + pow-1 : ∀ x → pow x 1 ≡ x + pow-* : ∀ x a b → pow x (a *ℤ b) ≡ pow (pow x a) b + pow-unit : ∀ n → pow unit n ≡ unit + pow-inverse : ∀ x n → pow (x ⁻¹) n ≡ pow x n ⁻¹ +``` + + +```agda + pow-sucl x a = + pow x (sucℤ a) ≡˘⟨ ap (pow x) (+ℤ-onel a) ⟩ + pow x (1 +ℤ a) ≡⟨ pow-+ x 1 a ⟩ + pow x 1 ⋆ pow x a ≡⟨ ap (_⋆ pow x a) (pow-1 x) ⟩ + x ⋆ pow x a ∎ + + pow-0 x = refl + + pow-1 x = idl + + pow-* x a = ℤ.induction (ap (pow x) (*ℤ-zeror a)) λ b → + pow x (a *ℤ b) ≡ pow (pow x a) b ≃⟨ _ , equiv→cancellable (⋆-equivr _) ⟩ + pow x (a *ℤ b) ⋆ pow x a ≡ pow (pow x a) b ⋆ pow x a ≃⟨ ∙-pre-equiv (pow-+ x (a *ℤ b) a) ⟩ + pow x (a *ℤ b +ℤ a) ≡ pow (pow x a) b ⋆ pow x a ≃⟨ ∙-pre-equiv (ap (pow x) (*ℤ-sucr a b)) ⟩ + pow x (a *ℤ sucℤ b) ≡ pow (pow x a) b ⋆ pow x a ≃⟨ ∙-post-equiv (sym (pow-sucr (pow x a) b)) ⟩ + pow x (a *ℤ sucℤ b) ≡ pow (pow x a) (sucℤ b) ≃∎ + + pow-unit = ℤ.induction refl (λ x → ∙-pre-equiv (pow-sucr unit x ∙ idr)) + + pow-inverse x = ℤ.induction (sym inv-unit) λ n → + pow (x ⁻¹) n ≡ pow x n ⁻¹ ≃⟨ _ , equiv→cancellable (⋆-equivr (x ⁻¹)) ⟩ + pow (x ⁻¹) n ⋆ x ⁻¹ ≡ pow x n ⁻¹ ⋆ x ⁻¹ ≃⟨ ∙-pre-equiv (pow-sucr (x ⁻¹) n) ⟩ + pow (x ⁻¹) (sucℤ n) ≡ pow x n ⁻¹ ⋆ x ⁻¹ ≃⟨ ∙-post-equiv (sym inv-comm) ⟩ + pow (x ⁻¹) (sucℤ n) ≡ (x ⋆ pow x n) ⁻¹ ≃⟨ ∙-post-equiv (ap _⁻¹ (sym (pow-sucl x n))) ⟩ + pow (x ⁻¹) (sucℤ n) ≡ pow x (sucℤ n) ⁻¹ ≃∎ +``` +
+ +Finally, the properties above imply that $\ZZ$ is the free group on +one generator. + +```agda +ℤ-free : Free-object Grp↪Sets (el! ⊤) +ℤ-free .Free-object.free = LiftGroup lzero ℤ +ℤ-free .Free-object.unit _ = 1 +ℤ-free .Free-object.fold {G} x = pow-hom G (x _) +ℤ-free .Free-object.commute {G} {x} = ext λ _ → pow-1 G (x _) +ℤ-free .Free-object.unique {G} {x} g comm = + pow-unique G (x _) g (unext comm _) +``` + +::: note +While the notation $x^n$ for `pow`{.Agda} makes sense in +multiplicative notation, we would instead write $n \cdot x$ in additive +notation. In fact, we can show that $n \cdot x$ coincides with the +multiplication $n \times x$ in the group of integers itself. + +```agda +pow-ℤ : ∀ a b → pow ℤ a b ≡ a *ℤ b +pow-ℤ a = ℤ.induction (sym (*ℤ-zeror a)) λ b → + pow ℤ a b ≡ a *ℤ b ≃⟨ ap (_+ℤ a) , equiv→cancellable (Group-on.⋆-equivr (ℤ .snd) a) ⟩ + pow ℤ a b +ℤ a ≡ a *ℤ b +ℤ a ≃⟨ ∙-pre-equiv (pow-sucr ℤ a b) ⟩ + pow ℤ a (sucℤ b) ≡ a *ℤ b +ℤ a ≃⟨ ∙-post-equiv (sym (*ℤ-sucr a b)) ⟩ + pow ℤ a (sucℤ b) ≡ a *ℤ sucℤ b ≃∎ +``` +::: + +# Order of an element {defines="order-of-an-element"} + + + +We define the **order** of an element $x : G$ of a group $G$ as the +minimal *positive*^[Without this requirement, every element would +trivially have order $0$!] integer $n$ such that $x^n = 1$, if it exists. + +In particular, if $x^n = 1$, then we have that the order of $x$ divides +$n$. We define this notion first in the code, then use it to define the +order of $x$. + +```agda + order-divides : ⌞ G ⌟ → Nat → Type ℓ + order-divides x n = pow G x (pos n) ≡ unit + + has-finite-order : ⌞ G ⌟ → Type ℓ + has-finite-order x = minimal-solution λ n → + Positive n × order-divides x n + + order : (x : ⌞ G ⌟) → has-finite-order x → Nat + order x (n , _) = n +``` + +::: {.definition #torsion} +An element $x$ with finite order is also called a **torsion element**. +A group all of whose elements are torsion is called a **torsion group**, +while a group whose only torsion element is the unit is called +**torsion-free**. +::: + +```agda + is-torsion-group : Type ℓ + is-torsion-group = ∀ g → has-finite-order g + + is-torsion-free : Type ℓ + is-torsion-free = ∀ g → has-finite-order g → g ≡ unit +``` + +::: note +Our definition of torsion groups requires being able to compute +the order of every element of the group. There is a weaker definition +that only requires every element $x$ to have *some* $n$ such that +$x^n = 1$; the two definitions are equivalent if the group is +[[discrete]], since [[$\NN$ is well-ordered|N is well-ordered]]. +::: + +We quickly note that $\ZZ$ is torsion-free, since multiplication by +a nonzero integer is injective. + +```agda +ℤ-torsion-free : is-torsion-free ℤ +ℤ-torsion-free n (k , (k>0 , nk≡0) , _) = *ℤ-injectiver (pos k) n 0 + (λ p → <-not-equal k>0 (pos-injective (sym p))) + (sym (pow-ℤ n (pos k)) ∙ nk≡0) +``` diff --git a/src/Algebra/Group/Instances/Symmetric.lagda.md b/src/Algebra/Group/Instances/Symmetric.lagda.md new file mode 100644 index 000000000..a3282da7d --- /dev/null +++ b/src/Algebra/Group/Instances/Symmetric.lagda.md @@ -0,0 +1,66 @@ + + +```agda +module Algebra.Group.Instances.Symmetric where +``` + +# Symmetric groups {defines="symmetric-group"} + +If $X$ is a set, then the type of all bijections $X \simeq X$ is also a +set, and it forms the carrier for a [[group]]: The _symmetric group_ on $X$. + +```agda +Sym : ∀ {ℓ} (X : Set ℓ) → Group-on (∣ X ∣ ≃ ∣ X ∣) +Sym X = to-group-on group-str where + open make-group + group-str : make-group (∣ X ∣ ≃ ∣ X ∣) + group-str .mul g f = f ∙e g +``` + +The group operation is `composition of equivalences`{.Agda ident=∙e}; +The identity element is `the identity equivalence`{.Agda ident=id-equiv}. + +```agda + group-str .unit = id , id-equiv +``` + +This type is a set because $X \to X$ is a set (because $X$ is a set by +assumption), and `being an equivalence is a proposition`{.Agdaa +ident=is-equiv-is-prop}. + +```agda + group-str .group-is-set = hlevel 2 +``` + +The associativity and identity laws hold definitionally. + +```agda + group-str .assoc _ _ _ = trivial! + group-str .idl _ = trivial! +``` + +The inverse is given by `the inverse equivalence`{.Agda ident=_e⁻¹}, and +the inverse equations hold by the fact that the inverse of an +equivalence is both a section and a retraction. + +```agda + group-str .inv = _e⁻¹ + group-str .invl (f , eqv) = ext (equiv→unit eqv) +``` + +We write $S_n$ for the symmetric group on the [[standard finite set]] +with $n$ elements. + +```agda +S : Nat → Group lzero +S n = el! (Fin n ≃ Fin n) , Sym (el! (Fin n)) +``` diff --git a/src/Algebra/Group/Semidirect.lagda.md b/src/Algebra/Group/Semidirect.lagda.md new file mode 100644 index 000000000..97e1c7e2d --- /dev/null +++ b/src/Algebra/Group/Semidirect.lagda.md @@ -0,0 +1,133 @@ + + +```agda +module Algebra.Group.Semidirect {ℓ} where +``` + +# Semidirect products of groups {defines="semidirect-product"} + +Given a [[group]] $A$ [[acting|group action]] on another group $B$ (that is, +equipped with a [[group homomorphism]] $\varphi : A \to \rm{Aut}(B)$ from $A$ to +the [[automorphism group]] of $B$), we define the **semidirect product** +$A \ltimes B$ as a variant of the [[direct product|direct product of groups]] $A \times B$ in which +the $A$ component influences the multiplication of the $B$ component. + +To motivate the concept, let's consider some examples: + +- The [[dihedral group]] of order $2n$, $D_n$, which is the group of +symmetries of a regular $n$-gon, can be defined as the semidirect product +$\ZZ/2\ZZ \ltimes \ZZ/n\ZZ$, where the action of $\ZZ/2\ZZ$ sends $1$ +to the [[negation automorphism]] of $\ZZ/n\ZZ$. Intuitively, a symmetry +of the regular $n$-gon has two components: a reflection (captured by the +cyclic group $\ZZ/2\ZZ$) and a rotation (captured by the cyclic group $\ZZ/n\ZZ$); +but when combining two symmetries, a reflection will swap the direction +of further rotations. This is exactly captured by the semidirect product. +- The group of isometries of the square lattice $\ZZ^2$ (considered +as a subspace of $\RR^2$), which we will refer to by its IUC notation +$p4m$, consists of reflections, rotations and translations. It can thus +be defined as a further semidirect product of the [[dihedral group]] +$D_4$ acting on the group of translations $\ZZ^2$ in the obvious way. +- Leaving the discrete world, the group of all isometries of the plane +$\RR^2$ can be seen analogously as the semidirect product of the +orthogonal group $O(2)$ acting on the translation group $\RR^2$, where +$O(2)$ is itself the semidirect product of $\ZZ/2\ZZ$ acting on the +*special* orthogonal group $SO(2)$, also known as the circle group. + + + +We are now ready to define $A \ltimes B$. The underlying set is the +cartesian product $A \times B$; the unit is the pair of units +$(1_A, 1_B)$. The interesting part of the definition is the multiplication: +we set $(a, b) \star (a', b') = (a \star_A a', \varphi_{a'}(b) \star_B b')$. + +::: note +Since our [[group actions]] are *right* actions, what we define here +is the *right* semidirect product. The version that works with *left* +actions instead would be $(a \star_A a', b \star_B \varphi_{a}(b'))$. +::: + +```agda + Semidirect-product : Group ℓ + Semidirect-product = to-group A⋉B where + A⋉B : make-group (⌞ A ⌟ × ⌞ B ⌟) + A⋉B .group-is-set = hlevel 2 + A⋉B .unit = A.unit , B.unit + A⋉B .mul (a , b) (a' , b') = a A.⋆ a' , (φ # a' # b) B.⋆ b' +``` + +The definition of the inverses is then forced, and we easily check that +the group axioms are verified. + +```agda + A⋉B .inv (a , b) = a A.⁻¹ , (φ # (a A.⁻¹) # b) B.⁻¹ + A⋉B .assoc (ax , bx) (ay , by) (az , bz) = Σ-pathp A.associative $ + ⌜ φ # (ay A.⋆ az) # bx ⌝ B.⋆ φ # az # by B.⋆ bz ≡⟨ ap! (unext (φ .preserves .pres-⋆ _ _) _) ⟩ + φ # az # (φ # ay # bx) B.⋆ φ # az # by B.⋆ bz ≡⟨ B.associative ⟩ + (φ # az # (φ # ay # bx) B.⋆ φ # az # by) B.⋆ bz ≡˘⟨ ap (B._⋆ bz) ((φ # az) .Groups.to .preserves .pres-⋆ _ _) ⟩ + φ # az # (φ # ay # bx B.⋆ by) B.⋆ bz ∎ + A⋉B .invl (a , b) = Σ-pathp A.inversel $ + φ # a # ⌜ (φ # (a A.⁻¹) # b) B.⁻¹ ⌝ B.⋆ b ≡˘⟨ ap¡ (pres-inv ((φ # _) .Groups.to .preserves)) ⟩ + φ # a # (φ # (a A.⁻¹) # (b B.⁻¹)) B.⋆ b ≡˘⟨ ap (B._⋆ b) (unext (pres-⋆ (φ .preserves) _ _) _) ⟩ + φ # ⌜ a A.⁻¹ A.⋆ a ⌝ # (b B.⁻¹) B.⋆ b ≡⟨ ap! A.inversel ⟩ + φ # A.unit # (b B.⁻¹) B.⋆ b ≡⟨ ap (B._⋆ b) (unext (pres-id (φ .preserves)) _) ⟩ + b B.⁻¹ B.⋆ b ≡⟨ B.inversel ⟩ + B.unit ∎ + A⋉B .idl (a , b) = Σ-pathp A.idl $ + φ # a # B.unit B.⋆ b ≡⟨ ap (B._⋆ b) (pres-id ((φ # a) .Groups.to .preserves)) ⟩ + B.unit B.⋆ b ≡⟨ B.idl ⟩ + b ∎ +``` + +The natural projection from $A \ltimes B$ to the first component $A$ is +a [[group homomorphism]], but the same cannot be said about the second +component (can you see why?). + +```agda + ⋉-proj : Groups.Hom Semidirect-product A + ⋉-proj .hom = fst + ⋉-proj .preserves .pres-⋆ _ _ = refl +``` + +When the action of $A$ on $B$ is [[trivial|trivial action]], i.e. $\varphi_a$ is the +identity automorphism for all $a : A$, it's easy to see that the +semidirect product $A \ltimes B$ agrees with the direct product +$A \times B$. + + + +```agda + Semidirect-trivial + : Semidirect-product A B (trivial-action A B) ≡ Direct-product A B + Semidirect-trivial = ∫-Path + (total-hom (λ x → x) (record { pres-⋆ = λ _ _ → refl })) + id-equiv +``` + +Another way to understand the semidirect product is to look at its +definition [[for concrete groups|semidirect product of concrete groups]], +which is remarkably simple. diff --git a/src/Algebra/Group/Subgroup.lagda.md b/src/Algebra/Group/Subgroup.lagda.md index 4138277d9..06d14ac30 100644 --- a/src/Algebra/Group/Subgroup.lagda.md +++ b/src/Algebra/Group/Subgroup.lagda.md @@ -387,6 +387,7 @@ f(yy\inv) = f(1) = 1$$. B.unit ∎ ``` +:::{.definition #normal-subgroup} It turns out that this last property is enough to pick out exactly the kernels amongst the representations of subgroups: If $H$ is closed under conjugation, then $H$ generates an equivalence relation on the set @@ -395,6 +396,7 @@ this equivalence relation with a group structure. The kernel of the quotient map $G \to G/H$ is then $H$. We call a predicate representing a kernel a **normal subgroup**, and we denote this in shorthand by $H \unlhd G$. +::: ```agda record normal-subgroup (G : Group ℓ) (H : ℙ ⌞ G ⌟) : Type ℓ where @@ -495,6 +497,7 @@ rather directly: Group-on-G/H .make-group.invl = elim! λ x → ap Coeq.inc inversel Group-on-G/H .make-group.idl = elim! λ x → ap Coeq.inc idl + infix 25 _/ᴳ_ _/ᴳ_ : Group _ _/ᴳ_ = to-group Group-on-G/H diff --git a/src/Algebra/Magma/Unital.lagda.md b/src/Algebra/Magma/Unital.lagda.md index a291c02aa..758b39ae3 100644 --- a/src/Algebra/Magma/Unital.lagda.md +++ b/src/Algebra/Magma/Unital.lagda.md @@ -142,7 +142,7 @@ record open Unital-magma≃ ``` -* One-sided identities +## One-sided identities Dropping either of the paths involved in a unital magma results in a right identity or a left identity. diff --git a/src/Algebra/Monoid/Category.lagda.md b/src/Algebra/Monoid/Category.lagda.md index a4d7a58cf..20958c37d 100644 --- a/src/Algebra/Monoid/Category.lagda.md +++ b/src/Algebra/Monoid/Category.lagda.md @@ -107,7 +107,7 @@ Mon↪Sets : ∀ {ℓ} → Functor (Monoids ℓ) (Sets ℓ) Mon↪Sets = Forget-structure (Monoid-structure _) ``` -## Free objects {defines=free-monoid} +## Free monoids {defines=free-monoid} We piece together some properties of `lists`{.Agda ident=List} to show that, if $A$ is a set, then $\rm{List}(A)$ is an object of diff --git a/src/Algebra/Ring.lagda.md b/src/Algebra/Ring.lagda.md index 5221dd515..d3ebbfe8d 100644 --- a/src/Algebra/Ring.lagda.md +++ b/src/Algebra/Ring.lagda.md @@ -2,7 +2,7 @@ ```agda open import Algebra.Group.Cat.Base open import Algebra.Semigroup -open import Algebra.Group.Ab hiding (ℤ) +open import Algebra.Group.Ab open import Algebra.Prelude open import Algebra.Monoid open import Algebra.Group diff --git a/src/Cat/Abelian/Endo.lagda.md b/src/Cat/Abelian/Endo.lagda.md index 758e0ed4d..836d9e33d 100644 --- a/src/Cat/Abelian/Endo.lagda.md +++ b/src/Cat/Abelian/Endo.lagda.md @@ -17,7 +17,7 @@ private module A = Ab-category A ``` --> -# Endomorphism rings +# Endomorphism rings {defines="endomorphism-ring"} Fix an [$\Ab$-category] $\cA$: It can be the category of [[abelian groups]] $\Ab$ itself, for example, or $R$-Mod for your favourite diff --git a/src/Cat/Bi/Instances/Relations.lagda.md b/src/Cat/Bi/Instances/Relations.lagda.md index f597e5e09..faaf13fd9 100644 --- a/src/Cat/Bi/Instances/Relations.lagda.md +++ b/src/Cat/Bi/Instances/Relations.lagda.md @@ -443,7 +443,7 @@ want to look at the formalisation. (image-pre-cover r[st].it _ β-cover) done : Im r[st].it Sub.≅ Im [rs]t.it - done = r[st]≅i Sub.∘Iso ([rs]t≅i Sub.Iso⁻¹) + done = [rs]t≅i Sub.Iso⁻¹ Sub.∘Iso r[st]≅i ``` --> diff --git a/src/Cat/Diagram/Initial.lagda.md b/src/Cat/Diagram/Initial.lagda.md index 7916c4ce9..12256894d 100644 --- a/src/Cat/Diagram/Initial.lagda.md +++ b/src/Cat/Diagram/Initial.lagda.md @@ -17,7 +17,7 @@ module _ {o h} (C : Precategory o h) where ``` --> -# Initial objects {defines="initial-object"} +# Initial objects {defines="initial-object initial"} An object $\bot$ of a category $\mathcal{C}$ is said to be **initial** if there exists a _unique_ map to any other object: @@ -25,7 +25,7 @@ if there exists a _unique_ map to any other object: ```agda is-initial : Ob → Type _ is-initial ob = ∀ x → is-contr (Hom ob x) - + record Initial : Type (o ⊔ h) where field bot : Ob @@ -38,13 +38,13 @@ contractible type, it is unique. ```agda ¡ : ∀ {x} → Hom bot x ¡ = has⊥ _ .centre - + ¡-unique : ∀ {x} (h : Hom bot x) → ¡ ≡ h ¡-unique = has⊥ _ .paths - + ¡-unique₂ : ∀ {x} (f g : Hom bot x) → f ≡ g ¡-unique₂ = is-contr→is-prop (has⊥ _) - + open Initial ``` @@ -85,7 +85,7 @@ a proposition: ⊥-is-prop : is-category C → is-prop Initial ⊥-is-prop ccat x1 x2 i .bot = Univalent.iso→path ccat (⊥-unique x1 x2) i - + ⊥-is-prop ccat x1 x2 i .has⊥ ob = is-prop→pathp (λ i → is-contr-is-prop @@ -107,7 +107,7 @@ This is an instance of the more general notion of [van Kampen colimits]. ```agda is-strict-initial : Initial → Type _ is-strict-initial i = ∀ x → (f : Hom x (i .bot)) → is-invertible f - + record Strict-initial : Type (o ⊔ h) where field initial : Initial @@ -125,7 +125,7 @@ Strictness is a property of, not structure on, an initial object. ```agda module _ {o h} {C : Precategory o h} where open Cat.Morphism C - private unquoteDecl eqv = declare-record-iso eqv (quote Initial) + private unquoteDecl eqv = declare-record-iso eqv (quote Initial) instance Extensional-Initial diff --git a/src/Cat/Diagram/Terminal.lagda.md b/src/Cat/Diagram/Terminal.lagda.md index c3438ac2d..ba9a405ab 100644 --- a/src/Cat/Diagram/Terminal.lagda.md +++ b/src/Cat/Diagram/Terminal.lagda.md @@ -20,7 +20,7 @@ module _ {o h} (C : Precategory o h) where ``` --> -# Terminal objects {defines=terminal-object} +# Terminal objects {defines="terminal-object terminal"} An object $\top$ of a category $\mathcal{C}$ is said to be **terminal** if it admits a _unique_ map from any other object: diff --git a/src/Cat/Displayed/Bifibration.lagda.md b/src/Cat/Displayed/Bifibration.lagda.md index 7383866fa..14e74c64b 100644 --- a/src/Cat/Displayed/Bifibration.lagda.md +++ b/src/Cat/Displayed/Bifibration.lagda.md @@ -96,7 +96,7 @@ module _ (bifib : is-bifibration) where → cobase-change f ⊣ base-change f cobase-change⊣base-change {x} {y} f = hom-natural-iso→adjoints $ - (opfibration→hom-iso opfibration f ni⁻¹) ∘ni fibration→hom-iso fibration f + fibration→hom-iso fibration f ∘ni opfibration→hom-iso opfibration f ni⁻¹ ``` In fact, if $\cE \liesover \cB$ is a cartesian fibration where every @@ -127,7 +127,7 @@ module _ (fib : Cartesian-fibration) where left-adjoint-base-change→opfibration L adj = cartesian+weak-opfibration→opfibration fib $ hom-iso→weak-opfibration L λ u → - fibration→hom-iso-from fib u ∘ni (adjunct-hom-iso-from (adj u) _ ni⁻¹) + adjunct-hom-iso-from (adj u) _ ni⁻¹ ∘ni fibration→hom-iso-from fib u ``` - Then, since the spaces of equivalences $\cC \cong \cD$ and isomorphisms $\cC \to \cD$ are both defined as the total space of a predicate on the same types, it suffices to show that the predicates @@ -317,3 +299,17 @@ the necessary paths for showing that $F_0$ is an equivalence of types. .is-iso.rinv x → dcat .to-path $ isoⁿ→iso (F∘F⁻¹≅Id eqv) _ .is-iso.linv x → sym $ ccat .to-path $ isoⁿ→iso (Id≅F⁻¹∘F eqv) _ ``` + + diff --git a/src/Cat/Functor/Hom/Duality.lagda.md b/src/Cat/Functor/Hom/Duality.lagda.md index 70e13bccc..3d2e10deb 100644 --- a/src/Cat/Functor/Hom/Duality.lagda.md +++ b/src/Cat/Functor/Hom/Duality.lagda.md @@ -39,13 +39,13 @@ corepresentable→co-representable : ∀ {F : Functor C (Sets ℓ)} → Corepresentation F → Representation {C = C ^op} F corepresentable→co-representable F-corep .rep = F-corep .corep -corepresentable→co-representable F-corep .represents = F-corep .corepresents - ∘ni path→iso (sym (Hom-into-op _)) +corepresentable→co-representable F-corep .represents = + path→iso (sym (Hom-into-op _)) ∘ni F-corep .corepresents co-representable→corepresentable : ∀ {F : Functor (C ^op) (Sets ℓ)} → Representation {C = C} F → Corepresentation F co-representable→corepresentable F-rep .corep = F-rep .rep -co-representable→corepresentable F-rep .corepresents = F-rep .represents - ∘ni path→iso (sym (Hom-from-op _)) +co-representable→corepresentable F-rep .corepresents = + path→iso (sym (Hom-from-op _)) ∘ni F-rep .represents ``` diff --git a/src/Cat/Functor/Hom/Representable.lagda.md b/src/Cat/Functor/Hom/Representable.lagda.md index a7b4df75b..e98d86ba9 100644 --- a/src/Cat/Functor/Hom/Representable.lagda.md +++ b/src/Cat/Functor/Hom/Representable.lagda.md @@ -91,7 +91,7 @@ representation-unique : {F : Functor (C ^op) (Sets κ)} (X Y : Representation F) representation-unique X Y = is-ff→essentially-injective {F = よ C} (よ-is-fully-faithful C) よX≅よY where よX≅よY : よ₀ C (X .rep) C^.≅ よ₀ C (Y .rep) - よX≅よY = (X .represents C^.Iso⁻¹) C^.∘Iso Y .represents + よX≅よY = Y .represents C^.∘Iso X .represents C^.Iso⁻¹ ``` Therefore, if $\cC$ is a [[univalent category]], then the type of @@ -274,7 +274,7 @@ corepresentation-unique X Y = (iso→co-iso (Cat[ C , Sets κ ]) ni) where ni : Hom-from C (Y .corep) ≅ⁿ Hom-from C (X .corep) - ni = (Y .corepresents ni⁻¹) ∘ni X .corepresents + ni = X .corepresents ∘ni Y .corepresents ni⁻¹ ``` diff --git a/src/Cat/Functor/Naturality.lagda.md b/src/Cat/Functor/Naturality.lagda.md index 47024aa2e..89395ea0d 100644 --- a/src/Cat/Functor/Naturality.lagda.md +++ b/src/Cat/Functor/Naturality.lagda.md @@ -58,13 +58,16 @@ isomorphism in a functor category. idni : ∀ {F} → F ≅ⁿ F idni = CD.id-iso - _∘ni_ : ∀ {F G H} → F ≅ⁿ G → G ≅ⁿ H → F ≅ⁿ H + _∘ni_ : ∀ {F G H} → G ≅ⁿ H → F ≅ⁿ G → F ≅ⁿ H _∘ni_ = CD._∘Iso_ + _∙ni_ : ∀ {F G H} → F ≅ⁿ G → G ≅ⁿ H → F ≅ⁿ H + f ∙ni g = g ∘ni f + _ni⁻¹ : ∀ {F G} → F ≅ⁿ G → G ≅ⁿ F _ni⁻¹ = CD._Iso⁻¹ - infixr 30 _∘ni_ + infixr 30 _∘ni_ _∙ni_ infix 31 _ni⁻¹ ≅ⁿ-pathp : ∀ {a c b d : Functor C D} (p : a ≡ c) (q : b ≡ d) {f : a ≅ⁿ b} {g : c ≅ⁿ d} diff --git a/src/Cat/Monoidal/Base.lagda.md b/src/Cat/Monoidal/Base.lagda.md index 5e083dd79..cacc5705c 100644 --- a/src/Cat/Monoidal/Base.lagda.md +++ b/src/Cat/Monoidal/Base.lagda.md @@ -176,8 +176,8 @@ children's drawing of a house, so that it fits on the page horizontally. → (A ▶ α→ B C D) ∘ α→ A (B ⊗ C) D ∘ (α→ A B C ◀ D) ≡ α→ A B (C ⊗ D) ∘ α→ (A ⊗ B) C D pentagon-α→ = inverse-unique refl refl - (▶.F-map-iso (α≅ Iso⁻¹) ∘Iso α≅ Iso⁻¹ ∘Iso ◀.F-map-iso (α≅ Iso⁻¹)) - (α≅ Iso⁻¹ ∘Iso α≅ Iso⁻¹) + (▶.F-map-iso (α≅ Iso⁻¹) ∙Iso α≅ Iso⁻¹ ∙Iso ◀.F-map-iso (α≅ Iso⁻¹)) + (α≅ Iso⁻¹ ∙Iso α≅ Iso⁻¹) (sym (assoc _ _ _) ∙ pentagon) ``` --> @@ -301,24 +301,24 @@ desired equation since $1 \otimes -$ is an equivalence. triangle-λ← {A} {B} = push-eqⁿ (unitor-l ni⁻¹) $ ▶.F-∘ _ _ ∙ ap to (Iso-prism base sq1 sq2 sq3) where - base : ◀.F-map-iso (α≅ Iso⁻¹) ∘Iso ◀.F-map-iso (◀.F-map-iso (ρ≅ Iso⁻¹)) + base : ◀.F-map-iso (α≅ Iso⁻¹) ∙Iso ◀.F-map-iso (◀.F-map-iso (ρ≅ Iso⁻¹)) ≡ ◀.F-map-iso (▶.F-map-iso (λ≅ Iso⁻¹)) base = ≅-path (◀.collapse triangle) - sq1 : ◀.F-map-iso (α≅ Iso⁻¹) ∘Iso α≅ ∘Iso α≅ ≡ α≅ ∘Iso ▶.F-map-iso α≅ + sq1 : ◀.F-map-iso (α≅ Iso⁻¹) ∙Iso α≅ ∙Iso α≅ ≡ α≅ ∙Iso ▶.F-map-iso α≅ sq1 = ≅-path (rswizzle (sym pentagon-α→ ∙ assoc _ _ _) (◀.annihilate (α≅ .invl))) - sq2 : ◀.F-map-iso (◀.F-map-iso (ρ≅ Iso⁻¹)) ∘Iso α≅ - ≡ (α≅ ∘Iso α≅) ∘Iso ▶.F-map-iso (λ≅ Iso⁻¹) + sq2 : ◀.F-map-iso (◀.F-map-iso (ρ≅ Iso⁻¹)) ∙Iso α≅ + ≡ (α≅ ∙Iso α≅) ∙Iso ▶.F-map-iso (λ≅ Iso⁻¹) sq2 = ≅-path $ α→ _ _ _ ∘ ((ρ← ⊗₁ id) ⊗₁ id) ≡⟨ associator .Isoⁿ.to .is-natural _ _ _ ⟩ (ρ← ⊗₁ ⌜ id ⊗₁ id ⌝) ∘ α→ _ _ _ ≡⟨ ap! ⊗.F-id ⟩ (ρ← ⊗₁ id) ∘ α→ _ _ _ ≡˘⟨ pulll triangle-α→ ⟩ (id ⊗₁ λ←) ∘ α→ _ _ _ ∘ α→ _ _ _ ∎ - sq3 : ◀.F-map-iso (▶.F-map-iso (λ≅ Iso⁻¹)) ∘Iso α≅ - ≡ α≅ ∘Iso ▶.F-map-iso (◀.F-map-iso (λ≅ Iso⁻¹)) + sq3 : ◀.F-map-iso (▶.F-map-iso (λ≅ Iso⁻¹)) ∙Iso α≅ + ≡ α≅ ∙Iso ▶.F-map-iso (◀.F-map-iso (λ≅ Iso⁻¹)) sq3 = ≅-path (associator .Isoⁿ.to .is-natural _ _ _) ``` diff --git a/src/Cat/Monoidal/Braided.lagda.md b/src/Cat/Monoidal/Braided.lagda.md index c86e1ccfe..1e71a3dd7 100644 --- a/src/Cat/Monoidal/Braided.lagda.md +++ b/src/Cat/Monoidal/Braided.lagda.md @@ -121,8 +121,8 @@ braiding. β←-α← : ∀ {A B C} → (β← ⊗₁ id) ∘ α← B A C ∘ (id ⊗₁ β←) ≡ α← A B C ∘ β← ∘ α← B C A β←-α← = inverse-unique refl refl - (◀.F-map-iso β≅ ∘Iso α≅ ∘Iso ▶.F-map-iso β≅) - (α≅ ∘Iso β≅ ∘Iso α≅) + (◀.F-map-iso β≅ ∙Iso α≅ ∙Iso ▶.F-map-iso β≅) + (α≅ ∙Iso β≅ ∙Iso α≅) (sym (assoc _ _ _) ·· braiding-α→ ·· assoc _ _ _) ``` --> diff --git a/src/Cat/Morphism.lagda.md b/src/Cat/Morphism.lagda.md index 492df4d99..d810b510a 100644 --- a/src/Cat/Morphism.lagda.md +++ b/src/Cat/Morphism.lagda.md @@ -420,33 +420,36 @@ id-iso .inverses .invr = idl id Isomorphism = _≅_ -Inverses-∘ : {f : Hom a b} {f⁻¹ : Hom b a} {g : Hom b c} {g⁻¹ : Hom c b} - → Inverses f f⁻¹ → Inverses g g⁻¹ → Inverses (g ∘ f) (f⁻¹ ∘ g⁻¹) +Inverses-∘ : {a b c : Ob} {f : Hom b c} {f⁻¹ : Hom c b} {g : Hom a b} {g⁻¹ : Hom b a} + → Inverses f f⁻¹ → Inverses g g⁻¹ → Inverses (f ∘ g) (g⁻¹ ∘ f⁻¹) Inverses-∘ {f = f} {f⁻¹} {g} {g⁻¹} finv ginv = record { invl = l ; invr = r } where module finv = Inverses finv module ginv = Inverses ginv abstract - l : (g ∘ f) ∘ f⁻¹ ∘ g⁻¹ ≡ id - l = (g ∘ f) ∘ f⁻¹ ∘ g⁻¹ ≡⟨ cat! C ⟩ - g ∘ (f ∘ f⁻¹) ∘ g⁻¹ ≡⟨ (λ i → g ∘ finv.invl i ∘ g⁻¹) ⟩ - g ∘ id ∘ g⁻¹ ≡⟨ cat! C ⟩ - g ∘ g⁻¹ ≡⟨ ginv.invl ⟩ + l : (f ∘ g) ∘ g⁻¹ ∘ f⁻¹ ≡ id + l = (f ∘ g) ∘ g⁻¹ ∘ f⁻¹ ≡⟨ cat! C ⟩ + f ∘ (g ∘ g⁻¹) ∘ f⁻¹ ≡⟨ (λ i → f ∘ ginv.invl i ∘ f⁻¹) ⟩ + f ∘ id ∘ f⁻¹ ≡⟨ cat! C ⟩ + f ∘ f⁻¹ ≡⟨ finv.invl ⟩ id ∎ - r : (f⁻¹ ∘ g⁻¹) ∘ g ∘ f ≡ id - r = (f⁻¹ ∘ g⁻¹) ∘ g ∘ f ≡⟨ cat! C ⟩ - f⁻¹ ∘ (g⁻¹ ∘ g) ∘ f ≡⟨ (λ i → f⁻¹ ∘ ginv.invr i ∘ f) ⟩ - f⁻¹ ∘ id ∘ f ≡⟨ cat! C ⟩ - f⁻¹ ∘ f ≡⟨ finv.invr ⟩ + r : (g⁻¹ ∘ f⁻¹) ∘ f ∘ g ≡ id + r = (g⁻¹ ∘ f⁻¹) ∘ f ∘ g ≡⟨ cat! C ⟩ + g⁻¹ ∘ (f⁻¹ ∘ f) ∘ g ≡⟨ (λ i → g⁻¹ ∘ finv.invr i ∘ g) ⟩ + g⁻¹ ∘ id ∘ g ≡⟨ cat! C ⟩ + g⁻¹ ∘ g ≡⟨ ginv.invr ⟩ id ∎ -_∘Iso_ : a ≅ b → b ≅ c → a ≅ c -(f ∘Iso g) .to = g .to ∘ f .to -(f ∘Iso g) .from = f .from ∘ g .from +_∘Iso_ : ∀ {a b c : Ob} → b ≅ c → a ≅ b → a ≅ c +(f ∘Iso g) .to = f .to ∘ g .to +(f ∘Iso g) .from = g .from ∘ f .from (f ∘Iso g) .inverses = Inverses-∘ (f .inverses) (g .inverses) -infixr 40 _∘Iso_ +_∙Iso_ : ∀ {a b c : Ob} → a ≅ b → b ≅ c → a ≅ c +f ∙Iso g = g ∘Iso f + +infixr 40 _∘Iso_ _∙Iso_ infixr 41 _Iso⁻¹ invertible-∘ @@ -455,7 +458,7 @@ invertible-∘ → is-invertible (f ∘ g) invertible-∘ f-inv g-inv = record { inv = g-inv.inv ∘ f-inv.inv - ; inverses = Inverses-∘ g-inv.inverses f-inv.inverses + ; inverses = Inverses-∘ f-inv.inverses g-inv.inverses } where module f-inv = is-invertible f-inv diff --git a/src/Cat/Morphism/Instances.agda b/src/Cat/Morphism/Instances.agda index 940afe8aa..0d234b73d 100644 --- a/src/Cat/Morphism/Instances.agda +++ b/src/Cat/Morphism/Instances.agda @@ -7,19 +7,20 @@ open Precategory module Cat.Morphism.Instances where -instance - H-Level-is-invertible - : ∀ {o ℓ} {C : Precategory o ℓ} {a b} {f : C .Hom a b} {n} → H-Level (is-invertible C f) (suc n) - H-Level-is-invertible {C = C} = prop-instance (is-invertible-is-prop C) - -unquoteDecl H-Level-↪ = declare-record-hlevel 2 H-Level-↪ (quote _↪_) -unquoteDecl H-Level-↠ = declare-record-hlevel 2 H-Level-↠ (quote _↠_) -unquoteDecl H-Level-Inverses = declare-record-hlevel 2 H-Level-Inverses (quote Inverses) -unquoteDecl H-Level-≅ = declare-record-hlevel 2 H-Level-≅ (quote _≅_) -unquoteDecl - H-Level-has-section = declare-record-hlevel 2 H-Level-has-section (quote has-section) -unquoteDecl - H-Level-has-retract = declare-record-hlevel 2 H-Level-has-retract (quote has-retract) +abstract + instance + H-Level-is-invertible + : ∀ {o ℓ} {C : Precategory o ℓ} {a b} {f : C .Hom a b} {n} → H-Level (is-invertible C f) (suc n) + H-Level-is-invertible {C = C} = prop-instance (is-invertible-is-prop C) + + unquoteDecl H-Level-↪ = declare-record-hlevel 2 H-Level-↪ (quote _↪_) + unquoteDecl H-Level-↠ = declare-record-hlevel 2 H-Level-↠ (quote _↠_) + unquoteDecl H-Level-Inverses = declare-record-hlevel 2 H-Level-Inverses (quote Inverses) + unquoteDecl H-Level-≅ = declare-record-hlevel 2 H-Level-≅ (quote _≅_) + unquoteDecl + H-Level-has-section = declare-record-hlevel 2 H-Level-has-section (quote has-section) + unquoteDecl + H-Level-has-retract = declare-record-hlevel 2 H-Level-has-retract (quote has-retract) module _ {o ℓ} {C : Precategory o ℓ} where instance Extensional-↪ : ∀ {a b ℓr} ⦃ _ : Extensional (C .Hom a b) ℓr ⦄ → Extensional (_↪_ C a b) ℓr @@ -30,3 +31,9 @@ module _ {o ℓ} {C : Precategory o ℓ} where instance Extensional-≅ : ∀ {a b ℓr} ⦃ _ : Extensional (C .Hom a b) ℓr ⦄ → Extensional (_≅_ C a b) ℓr Extensional-≅ ⦃ sa ⦄ = injection→extensional! (≅-pathp C refl refl) sa + + Funlike-≅ + : ∀ {ℓ' ℓ''} {A : Type ℓ'} {B : A → Type ℓ''} + → ∀ {a b} ⦃ i : Funlike (Hom C a b) A B ⦄ + → Funlike (_≅_ C a b) A B + Funlike-≅ .Funlike._#_ f x = f .to # x diff --git a/src/Cat/Reasoning.lagda.md b/src/Cat/Reasoning.lagda.md index b87406fe8..ab8391208 100644 --- a/src/Cat/Reasoning.lagda.md +++ b/src/Cat/Reasoning.lagda.md @@ -340,8 +340,8 @@ of isomorphisms. ```agda Iso-swapr : ∀ {a : x ≅ y} {b : y ≅ z} {c : x ≅ z} - → a ∘Iso b ≡ c - → a ≡ c ∘Iso (b Iso⁻¹) + → b ∘Iso a ≡ c + → a ≡ b Iso⁻¹ ∘Iso c Iso-swapr {a = a} {b = b} {c = c} p = ≅-path $ a .to ≡⟨ introl (b .invr) ⟩ (b .from ∘ b .to) ∘ a .to ≡⟨ pullr (ap to p) ⟩ @@ -349,8 +349,8 @@ Iso-swapr {a = a} {b = b} {c = c} p = ≅-path $ Iso-swapl : ∀ {a : x ≅ y} {b : y ≅ z} {c : x ≅ z} - → a ∘Iso b ≡ c - → b ≡ (a Iso⁻¹) ∘Iso c + → b ∘Iso a ≡ c + → b ≡ c ∘Iso a Iso⁻¹ Iso-swapl {a = a} {b = b} {c = c} p = ≅-path $ b .to ≡⟨ intror (a .invl) ⟩ b .to ∘ a .to ∘ a .from ≡⟨ pulll (ap to p) ⟩ @@ -384,11 +384,11 @@ bottom face. Iso-prism : ∀ {a : u ≅ v} {b : v ≅ w} {c : u ≅ w} → {d : u ≅ x} {e : v ≅ y} {f : w ≅ z} → {g : x ≅ y} {h : y ≅ z} {i : x ≅ z} - → a ∘Iso b ≡ c - → a ∘Iso e ≡ d ∘Iso g - → b ∘Iso f ≡ e ∘Iso h - → c ∘Iso f ≡ d ∘Iso i - → g ∘Iso h ≡ i + → b ∘Iso a ≡ c + → e ∘Iso a ≡ g ∘Iso d + → f ∘Iso b ≡ h ∘Iso e + → f ∘Iso c ≡ i ∘Iso d + → h ∘Iso g ≡ i Iso-prism {a = a} {b} {c} {d} {e} {f} {g} {h} {i} top left right front = ≅-path $ h .to ∘ g .to ≡⟨ ap₂ _∘_ (ap to (Iso-swapl (sym right))) (ap to (Iso-swapl (sym left)) ∙ sym (assoc _ _ _)) ⟩ diff --git a/src/Cat/Univalent.lagda.md b/src/Cat/Univalent.lagda.md index 3cddaea93..2f383d07e 100644 --- a/src/Cat/Univalent.lagda.md +++ b/src/Cat/Univalent.lagda.md @@ -176,6 +176,11 @@ paths in `Hom`{.Agda}-sets. ) p q + path→iso-∙ + : ∀ {A B C} (p : A ≡ B) (q : B ≡ C) + → path→iso (p ∙ q) ≡ path→iso p ∙Iso path→iso q + path→iso-∙ p q = ext (path→to-∙ p q) + path→to-sym : ∀ {A B} (p : A ≡ B) → path→iso p .from ≡ path→iso (sym p) .to path→to-sym = J (λ B p → path→iso p .from ≡ path→iso (sym p) .to) refl diff --git a/src/Cat/Univalent/Rezk/Universal.lagda.md b/src/Cat/Univalent/Rezk/Universal.lagda.md index 789084ec7..b0bfd7225 100644 --- a/src/Cat/Univalent/Rezk/Universal.lagda.md +++ b/src/Cat/Univalent/Rezk/Universal.lagda.md @@ -351,7 +351,7 @@ candidate over it. ```agda obj' : ∀ {b} → Essential-fibre H b → Obs b obj' (a₀ , h₀) .fst = F.₀ a₀ - obj' (a₀ , h₀) .snd .fst a h = F-map-iso F (H.iso.from (h B.∘Iso (h₀ B.Iso⁻¹))) + obj' (a₀ , h₀) .snd .fst a h = F-map-iso F (H.iso.from (h₀ B.Iso⁻¹ B.∘Iso h)) obj' (a₀ , h₀) .snd .snd (a , h) (a' , h') f p = F.collapse (H.ipushr p) Obs-is-prop : ∀ {b} (f : Essential-fibre H b) (c : Obs b) → obj' f ≡ c @@ -367,7 +367,7 @@ candidate over it. c = ckα .fst k = ckα .snd .fst α = ckα .snd .snd - c≅c' = (k a₀ h₀ C.Iso⁻¹) C.∘Iso k' a₀ h₀ + c≅c' = k' a₀ h₀ C.∘Iso k a₀ h₀ C.Iso⁻¹ ``` diff --git a/src/Data/Fin/Closure.lagda.md b/src/Data/Fin/Closure.lagda.md index 86945a535..126de3e2f 100644 --- a/src/Data/Fin/Closure.lagda.md +++ b/src/Data/Fin/Closure.lagda.md @@ -4,6 +4,7 @@ open import 1Lab.Prelude open import Data.Set.Coequaliser open import Data.Fin.Properties +open import Data.Nat.Prime open import Data.Fin.Base open import Data.Dec open import Data.Sum @@ -135,7 +136,7 @@ Finite-multiply {n} {m} = sum≡* (suc n) m = ap (m +_) (sum≡* n m) ``` -## Products +### Products Similarly to the case for sums, the cardinality of a dependent *product* of finite sets is the `product`{.Agda} of the cardinalities: @@ -158,6 +159,53 @@ Finite-product {suc n} B = Fin (B fzero * product n (B ∘ fsuc)) ≃∎ ``` +## Permutations + +We show that the set of permutations $[n] \simeq [n]$ is finite with +cardinality $n!$ (the `factorial`{.Agda} of $n$). + +We start by showing that $([n+1] \simeq [n+1]) \simeq [n+1] \times +([n] \simeq [n])$: a permutation of $[n+1]$ is determined by what happens +to $0$ and by the remaining permutation of $[n]$. + +```agda +Fin-permutations-suc + : ∀ n → (Fin (suc n) ≃ Fin (suc n)) ≃ (Fin (suc n) × (Fin n ≃ Fin n)) +Fin-permutations-suc n = to , is-iso→is-equiv is where + to : (Fin (suc n) ≃ Fin (suc n)) → Fin (suc n) × (Fin n ≃ Fin n) + to e .fst = e # 0 + to e .snd .fst i = avoid (e # 0) (e # (fsuc i)) λ p → + fzero≠fsuc (Equiv.injective e p) + to e .snd .snd = Fin-injection→equiv _ λ p → + fsuc-inj (Equiv.injective e (avoid-injective (e # 0) p)) + + is : is-iso to + is .is-iso.inv (n , e) .fst fzero = n + is .is-iso.inv (n , e) .fst (fsuc x) = skip n (e # x) + is .is-iso.inv (n , e) .snd = Fin-injection→equiv _ λ where + {fzero} {fzero} p → refl + {fzero} {fsuc y} p → absurd (skip-skips n _ (sym p)) + {fsuc x} {fzero} p → absurd (skip-skips n _ p) + {fsuc x} {fsuc y} p → ap fsuc (Equiv.injective e (skip-injective n _ _ p)) + is .is-iso.rinv (n , e) = Σ-pathp refl (ext λ i → avoid-skip n (e # i)) + is .is-iso.linv e = ext λ where + fzero → refl + (fsuc i) → skip-avoid (e # 0) (e # (fsuc i)) +``` + +We can now show that $([n] \simeq [n]) \simeq [n!]$ by induction. + +```agda +Fin-permutations : ∀ n → (Fin n ≃ Fin n) ≃ Fin (factorial n) +Fin-permutations zero = is-contr→≃ + (contr id≃ λ _ → ext λ ()) Finite-one-is-contr +Fin-permutations (suc n) = + Fin (suc n) ≃ Fin (suc n) ≃⟨ Fin-permutations-suc n ⟩ + Fin (suc n) × (Fin n ≃ Fin n) ≃⟨ Σ-ap-snd (λ _ → Fin-permutations n) ⟩ + Fin (suc n) × Fin (factorial n) ≃⟨ Finite-multiply ⟩ + Fin (factorial (suc n)) ≃∎ +``` + ## Decidable subsets Given a [[decidable]] predicate on $[n]$, we can compute $s$ such that diff --git a/src/Data/Fin/Finite.lagda.md b/src/Data/Fin/Finite.lagda.md index 3b296c199..ad90bcc1d 100644 --- a/src/Data/Fin/Finite.lagda.md +++ b/src/Data/Fin/Finite.lagda.md @@ -242,7 +242,9 @@ Finite-Σ {A = A} {P = P} ⦃ afin ⦄ ⦃ pfin ⦄ = ∥-∥-out! do Finite-⊥ = fin (inc (Finite-zero-is-initial e⁻¹)) Finite-⊤ = fin (inc (is-contr→≃⊤ Finite-one-is-contr e⁻¹)) -Finite-Bool = fin (inc (Iso→Equiv enum)) where + +Bool≃Fin2 : Bool ≃ Fin 2 +Bool≃Fin2 = Iso→Equiv enum where enum : Iso Bool (Fin 2) enum .fst false = 0 enum .fst true = 1 @@ -253,6 +255,8 @@ Finite-Bool = fin (inc (Iso→Equiv enum)) where enum .snd .is-iso.linv true = refl enum .snd .is-iso.linv false = refl +Finite-Bool = fin (inc Bool≃Fin2) + Finite-PathP = subst Finite (sym (PathP≡Path _ _ _)) (Discrete→Finite≡ Finite→Discrete) Finite-Lift = Finite-≃ (Lift-≃ e⁻¹) diff --git a/src/Data/Fin/Properties.lagda.md b/src/Data/Fin/Properties.lagda.md index e9e37fed9..fbdc8a807 100644 --- a/src/Data/Fin/Properties.lagda.md +++ b/src/Data/Fin/Properties.lagda.md @@ -178,6 +178,38 @@ avoid-injective {suc n} (fsuc i) {fzero} {fsuc k} p = absurd (fzero≠fsuc p) avoid-injective {suc n} (fsuc i) {fsuc j} {fzero} p = absurd (fzero≠fsuc (sym p)) avoid-injective {suc n} (fsuc i) {fsuc j} {fsuc k} p = ap fsuc (avoid-injective {n} i {j} {k} (fsuc-inj p)) + +skip-injective + : ∀ {n} (i : Fin (suc n)) (j k : Fin n) + → skip i j ≡ skip i k → j ≡ k +skip-injective fzero j k p = fsuc-inj p +skip-injective (fsuc i) fzero fzero p = refl +skip-injective (fsuc i) fzero (fsuc k) p = absurd (fzero≠fsuc p) +skip-injective (fsuc i) (fsuc j) fzero p = absurd (fzero≠fsuc (sym p)) +skip-injective (fsuc i) (fsuc j) (fsuc k) p = ap fsuc (skip-injective i j k (fsuc-inj p)) + +skip-skips + : ∀ {n} (i : Fin (suc n)) (j : Fin n) + → skip i j ≠ i +skip-skips fzero j p = fzero≠fsuc (sym p) +skip-skips (fsuc i) fzero p = fzero≠fsuc p +skip-skips (fsuc i) (fsuc j) p = skip-skips i j (fsuc-inj p) + +avoid-skip + : ∀ {n} (i : Fin (suc n)) (j : Fin n) {neq : i ≠ skip i j} + → avoid i (skip i j) neq ≡ j +avoid-skip fzero fzero = refl +avoid-skip fzero (fsuc j) = refl +avoid-skip (fsuc i) fzero = refl +avoid-skip (fsuc i) (fsuc j) = ap fsuc (avoid-skip i j) + +skip-avoid + : ∀ {n} (i : Fin (suc n)) (j : Fin (suc n)) {i≠j : i ≠ j} + → skip i (avoid i j i≠j) ≡ j +skip-avoid fzero fzero {i≠j} = absurd (i≠j refl) +skip-avoid {suc n} fzero (fsuc j) = refl +skip-avoid {suc n} (fsuc i) fzero = refl +skip-avoid {suc n} (fsuc i) (fsuc j) = ap fsuc (skip-avoid i j) ``` ## Iterated products and sums {defines="iterated-products"} diff --git a/src/Data/Int/Base.lagda.md b/src/Data/Int/Base.lagda.md index 183856700..157aa2073 100644 --- a/src/Data/Int/Base.lagda.md +++ b/src/Data/Int/Base.lagda.md @@ -15,14 +15,14 @@ open import Data.Nat.Base module Data.Int.Base where ``` -# The integers {defines=integers} +# The integers {defines=integer} The familiar [[set]] of integers, in addition to its intuitive characterisation in terms of positive and negative numbers, can be specified by a wide variety of universal properties: - The set of integers, with the operation of addition, is the [[free -group]] on one generator. +group]] on one generator (see the [[group of integers]]). - The set of integers, with the operations of addition and multiplication, is the [[initial ring]]. @@ -255,6 +255,13 @@ natural numbers. If the numbers have mismatched sign, then the addition function is actually computing a difference, and we already know how to compute differences. + + ```agda _+ℤ_ : Int → Int → Int pos x +ℤ pos y = pos (x + y) diff --git a/src/Data/Int/DivMod.lagda.md b/src/Data/Int/DivMod.lagda.md new file mode 100644 index 000000000..a6fa2f9ff --- /dev/null +++ b/src/Data/Int/DivMod.lagda.md @@ -0,0 +1,231 @@ + + +```agda +module Data.Int.DivMod where +``` + + + +# Integer division {defines="integer-division"} + +We define the division $a/b$ where $a$ is an [[integer]] and $b$ is +a positive [[natural number]]. + +```agda +record DivModℤ (a : Int) (b : Nat) : Type where + constructor divmodℤ + + field + quotℤ : Int + remℤ : Nat + .quotient : a ≡ quotℤ *ℤ pos b +ℤ pos remℤ + .smaller : remℤ < b + +open DivModℤ +``` + +In order to compute integer division, we reuse [[natural division]]. +We proceed by cases: if $a$ is positive, we can simply compute the +division $a/b$ as natural numbers. + +```agda +divide-posℤ : ∀ a b → .⦃ _ : Positive b ⦄ → DivModℤ a b +divide-posℤ (pos a) b using divmod q r p s ← divide-pos a b + = divmodℤ (pos q) r (ap pos p ∙ sym (ap (_+ℤ pos r) (assign-pos _))) s +``` + +If $a$ is negative, we can compute $(-a)/b$ as natural numbers and +distinguish two cases: if the remainder is $0$, then $b$ divides $a$, +so we simply negate the quotient and set the remainder to $0$. + +```agda +divide-posℤ (negsuc a) b@(suc b') with divide-pos (suc a) b +... | divmod q zero p s = divmodℤ (assign neg q) 0 + (ap (assign neg) p ∙ assign-+ (q * b) 0 ∙ ap (_+ℤ 0) (assign-*l q b)) + (s≤s 0≤x) +``` + +If the remainder is non-zero, we have to correct for the fact that +integer division as we've defined it rounds towards $-\infty$, while +natural division rounds towards $0$ (hence towards $+\infty$, for +negative numbers), so we take $-((-a)/b) - 1$ as the quotient and +$b - (-a)\%b$ as the remainder. + +```agda +... | divmod q (suc r) p s = divmodℤ (negsuc q) (b - suc r) + (ap (assign neg) p ∙ lemma) + (s≤s (monus-≤ b' r)) +``` + +
+One annoying lemma later, we are done. + +```agda + where + lemma : assign neg (q * b + suc r) ≡ pos (b' - r) +ℤ negsuc (b' + q * b) + lemma = + assign neg (q * b + suc r) ≡⟨ ap (assign neg) (+-commutative (q * b) _) ⟩ + negsuc (r + q * b) ≡˘⟨ negℤ-+ℤ-negsuc r (q * b) ⟩ + negℤ (pos r) +ℤ negsuc (q * b) ≡⟨ ap (_+ℤ negsuc (q * b)) (ℤ.insertl {h = negℤ (pos b')} (+ℤ-invl (pos b')) {f = negℤ (pos r)}) ⟩ + ⌜ negℤ (pos b') ⌝ +ℤ (pos b' -ℤ pos r) +ℤ negsuc (q * b) ≡˘⟨ ap¡ (assign-neg b') ⟩ + assign neg b' +ℤ (pos b' -ℤ pos r) +ℤ negsuc (q * b) ≡⟨ ap (_+ℤ negsuc (q * b)) (+ℤ-commutative (assign neg b') (pos b' -ℤ pos r)) ⟩ + (pos b' -ℤ pos r) +ℤ assign neg b' +ℤ negsuc (q * b) ≡˘⟨ +ℤ-assoc (pos b' -ℤ pos r) _ _ ⟩ + (pos b' -ℤ pos r) +ℤ (assign neg b' +ℤ negsuc (q * b)) ≡⟨ ap₂ _+ℤ_ (pos-pos b' r) (ap (_+ℤ negsuc (q * b)) (assign-neg b')) ⟩ + (b' ℕ- r) +ℤ (negℤ (pos b') +ℤ negsuc (q * b)) ≡⟨ ap ((b' ℕ- r) +ℤ_) (negℤ-+ℤ-negsuc b' (q * b)) ⟩ + (b' ℕ- r) +ℤ negsuc (b' + q * b) ≡⟨ ap₂ _+ℤ_ (nat-diff-monus b' r (≤-peel (<-weaken (recover s)))) refl ⟩ + pos (b' - r) +ℤ negsuc (b' + q * b) ∎ +``` +
+ +```agda +_/ℤ_ : Int → (b : Nat) → .⦃ _ : Positive b ⦄ → Int +a /ℤ b = divide-posℤ a b .quotℤ + +_%ℤ_ : Int → (b : Nat) → .⦃ _ : Positive b ⦄ → Nat +a %ℤ b = divide-posℤ a b .remℤ + +is-divmodℤ : ∀ x y → .⦃ _ : Positive y ⦄ → x ≡ (x /ℤ y) *ℤ pos y +ℤ pos (x %ℤ y) +is-divmodℤ x (suc y) with divide-posℤ x (suc y) +... | divmodℤ q r α β = recover α + +x%ℤy + +```agda +module Data.Int.Divisible where +``` + +# Divisibility of integers {defines="divisibility-of-integers"} + +We define what it means for an [[integer]] to be [[divisible]] by a +natural number, as a straightforward generalisation of the case for +natural numbers. + +```agda +_∣ℤ_ : Nat → Int → Type +zero ∣ℤ y = y ≡ 0 +suc x ∣ℤ y = fibre (_*ℤ possuc x) y + +infix 7 _∣ℤ_ + +∣ℤ-is-prop : ∀ x y → is-prop (x ∣ℤ y) +∣ℤ-is-prop zero y n k = prop! +∣ℤ-is-prop (suc x) y (n , p) (n' , q) = Σ-prop-path! (*ℤ-injectiver-possuc x n n' (p ∙ sym q)) + +∣ℤ-is-truncation : ∀ {x y} → (x ∣ℤ y) ≃ ∥ fibre (_*ℤ pos x) y ∥ +∣ℤ-is-truncation {zero} {y} = + prop-ext! + (λ p → inc (y , *ℤ-zeror y ∙ sym p)) + (rec! λ x p → sym p ∙ *ℤ-zeror x ) +∣ℤ-is-truncation {suc x} {y} = Equiv.to is-prop≃equiv∥-∥ (∣ℤ-is-prop (suc x) y) + +∣ℤ→fibre : ∀ {x y} → x ∣ℤ y → fibre (_*ℤ pos x) y +∣ℤ→fibre {zero} wit = 0 , sym wit +∣ℤ→fibre {suc x} wit = wit + +fibre→∣ℤ : ∀ {x y} → fibre (_*ℤ pos x) y → x ∣ℤ y +fibre→∣ℤ f = Equiv.from ∣ℤ-is-truncation (inc f) + +dividesℤ : ∀ {x y} (q : Int) → q *ℤ pos x ≡ y → x ∣ℤ y +dividesℤ x p = fibre→∣ℤ (x , p) + +∣ℤ-zero : ∀ {x} → x ∣ℤ 0 +∣ℤ-zero = dividesℤ 0 refl + +∣ℤ-one : ∀ {x} → 1 ∣ℤ x +∣ℤ-one {x} = dividesℤ x (*ℤ-oner x) + +m∣ℤsn→m≤sn : ∀ {x y} → ¬ (y ≡ 0) → x ∣ℤ y → x Nat.≤ abs y +m∣ℤsn→m≤sn {x} {y} nz p with ∣ℤ→fibre p +... | posz , p = absurd (nz (sym p)) +... | possuc k , p = difference→≤ (k * x) + (ap abs (sym (assign-pos (x + k * x))) ∙ ap abs p) +... | negsuc k , p = difference→≤ (k * x) + (sym (abs-negℤ (pos (x + k * x))) ∙ ap abs (sym (assign-neg (x + k * x))) ∙ ap abs p) + +∣ℤ-+ : ∀ {k n m} → k ∣ℤ n → k ∣ℤ m → k ∣ℤ (n +ℤ m) +∣ℤ-+ {zero} {n} {m} p q = ap₂ _+ℤ_ p q +∣ℤ-+ {suc k} (x , p) (y , q) = x +ℤ y , *ℤ-distribr (possuc k) x y ∙ ap₂ _+ℤ_ p q + +∣ℤ-negℤ : ∀ {k n} → k ∣ℤ n → k ∣ℤ negℤ n +∣ℤ-negℤ {zero} d = ap negℤ d +∣ℤ-negℤ {suc k} (x , p) = negℤ x , *ℤ-negl x _ ∙ ap negℤ p +``` diff --git a/src/Data/Int/Properties.lagda.md b/src/Data/Int/Properties.lagda.md index 07ce298fb..ced6cc3c2 100644 --- a/src/Data/Int/Properties.lagda.md +++ b/src/Data/Int/Properties.lagda.md @@ -4,6 +4,7 @@ open import 1Lab.Prelude open import Data.Nat.Properties open import Data.Nat.Solver +open import Data.Nat.Order open import Data.Int.Base open import Data.Nat.Base ``` @@ -32,6 +33,11 @@ no further comments. nat-diff-zero zero = refl nat-diff-zero (suc x) = nat-diff-zero x + nat-diff-positive : ∀ a b → a ℕ- b ≡ 0 → a ≡ b + nat-diff-positive a zero p = pos-injective p + nat-diff-positive zero (suc b) p = absurd (negsuc≠pos p) + nat-diff-positive (suc a) (suc b) p = ap suc (nat-diff-positive a b p) + sucℤ-nat-diff : ∀ x y → sucℤ (x ℕ- y) ≡ suc x ℕ- y sucℤ-nat-diff zero zero = refl sucℤ-nat-diff zero (suc zero) = refl @@ -44,6 +50,16 @@ no further comments. predℤ-nat-diff zero (suc y) = refl predℤ-nat-diff (suc x) zero = refl predℤ-nat-diff (suc x) (suc y) = predℤ-nat-diff x y + + nat-diff-monus : ∀ x y → y ≤ x → x ℕ- y ≡ pos (x - y) + nat-diff-monus x y 0≤x = refl + nat-diff-monus (suc x) (suc zero) (s≤s y≤x) = refl + nat-diff-monus (suc x) (suc (suc y)) (s≤s y≤x) = nat-diff-monus x (suc y) y≤x + + nat-diff-bounded : ∀ a b c → a ≤ c → b ≤ c → abs (a ℕ- b) ≤ c + nat-diff-bounded a zero c ac bc = ac + nat-diff-bounded zero (suc b) c ac bc = bc + nat-diff-bounded (suc a) (suc b) c ac bc = nat-diff-bounded a b c (<-weaken ac) (<-weaken bc) ``` ## Negations @@ -92,6 +108,19 @@ no further comments. negℤ-distrib-min (negsuc x) (negsuc y) = refl ``` +## Absolute value + +```agda + abs-positive : ∀ x → abs x ≡ 0 → x ≡ 0 + abs-positive (pos x) p = ap pos p + abs-positive (negsuc x) p = absurd (suc≠zero p) + + abs-negℤ : ∀ z → abs (negℤ z) ≡ abs z + abs-negℤ (posz) = refl + abs-negℤ (possuc _) = refl + abs-negℤ (negsuc _) = refl +``` + ## Rotations ```agda @@ -201,6 +230,12 @@ no further comments. +ℤ-predr : ∀ x y → x +ℤ predℤ y ≡ predℤ (x +ℤ y) +ℤ-predr x y = +ℤ-commutative x (predℤ y) ·· +ℤ-predl y x ·· ap predℤ (+ℤ-commutative y x) + +ℤ-onel : ∀ x → 1 +ℤ x ≡ sucℤ x + +ℤ-onel x = +ℤ-sucl 0 x ∙ ap sucℤ (+ℤ-zerol x) + + +ℤ-oner : ∀ x → x +ℤ 1 ≡ sucℤ x + +ℤ-oner x = +ℤ-sucr x 0 ∙ ap sucℤ (+ℤ-zeror x) + +ℤ-injectiver : ∀ k x y → k +ℤ x ≡ k +ℤ y → x ≡ y +ℤ-injectiver k x y p = sym (+ℤ-zerol x) @@ -220,6 +255,39 @@ no further comments. ap negℤ (rot-is-add x y) ·· negℤ-distrib-rot x y ·· sym (rot-is-add (negℤ x) (negℤ y)) + + negℤ-+ℤ-negsuc : ∀ a b → negℤ (pos a) +ℤ negsuc b ≡ negsuc (a + b) + negℤ-+ℤ-negsuc a b = + +ℤ-commutative (negℤ (pos a)) (negsuc b) + ·· sym (negℤ-distrib (possuc b) (pos a)) + ·· ap negsuc (+-commutative b a) + + pos-pos : ∀ a b → pos a -ℤ pos b ≡ a ℕ- b + pos-pos a zero = +ℤ-zeror _ + pos-pos a (suc b) = refl + + -ℤ-swapl : ∀ a b c → a +ℤ b ≡ c → a ≡ c -ℤ b + -ℤ-swapl a b c p = + sym (+ℤ-zeror _) + ∙ ap (a +ℤ_) (sym (+ℤ-invr b)) + ·· +ℤ-assoc a _ _ + ·· ap (_+ℤ negℤ b) p + + private + distrib-lemma + : ∀ x y z w → (x +ℤ y) +ℤ (z +ℤ w) ≡ (x +ℤ z) +ℤ (y +ℤ w) + distrib-lemma x y z w = + +ℤ-assoc (x +ℤ y) z w + ·· ap (_+ℤ w) (sym (+ℤ-assoc x y z) ·· ap (x +ℤ_) (+ℤ-commutative y z) ·· +ℤ-assoc x z y) + ·· sym (+ℤ-assoc (x +ℤ z) y w) + + -ℤ-cancelr : ∀ k x y → (x +ℤ k) -ℤ (y +ℤ k) ≡ x -ℤ y + -ℤ-cancelr k x y = + (x +ℤ k) -ℤ (y +ℤ k) ≡⟨ ap ((x +ℤ k) +ℤ_) (negℤ-distrib y k) ⟩ + (x +ℤ k) +ℤ (negℤ y -ℤ k) ≡⟨ distrib-lemma x k (negℤ y) (negℤ k) ⟩ + (x -ℤ y) +ℤ (k -ℤ k) ≡⟨ ap ((x -ℤ y) +ℤ_) (+ℤ-invr k) ⟩ + (x -ℤ y) +ℤ 0 ≡⟨ +ℤ-zeror (x -ℤ y) ⟩ + x -ℤ y ∎ ``` ## Multiplication @@ -229,6 +297,32 @@ no further comments. assign-pos zero = refl assign-pos (suc x) = refl + assign-neg : ∀ x → assign neg x ≡ negℤ (pos x) + assign-neg zero = refl + assign-neg (suc x) = refl + + neg-sign : Sign → Sign + neg-sign pos = neg + neg-sign neg = pos + + neg-assign : ∀ {s} x → assign (neg-sign s) x ≡ negℤ (assign s x) + neg-assign {pos} zero = refl + neg-assign {pos} (suc x) = refl + neg-assign {neg} zero = refl + neg-assign {neg} (suc x) = refl + + assign-+ : ∀ {s} x y → assign s (x + y) ≡ assign s x +ℤ assign s y + assign-+ {pos} zero y = sym (+ℤ-zerol _) + assign-+ {pos} (suc x) zero = refl + assign-+ {pos} (suc x) (suc y) = refl + assign-+ {neg} zero y = sym (+ℤ-zerol _) + assign-+ {neg} (suc x) zero = ap negsuc (+-zeror _) + assign-+ {neg} (suc x) (suc y) = ap negsuc (+-sucr _ _) + + possuc≠assign-neg : ∀ {x y} → possuc x ≠ assign neg y + possuc≠assign-neg {x} {zero} p = suc≠zero (pos-injective p) + possuc≠assign-neg {x} {suc y} p = pos≠negsuc p + *ℤ-onel : ∀ x → 1 *ℤ x ≡ x *ℤ-onel (pos x) = ap (assign pos) (+-zeror x) ∙ assign-pos x *ℤ-onel (negsuc x) = ap negsuc (+-zeror x) @@ -245,6 +339,21 @@ no further comments. *ℤ-zeror (pos x) = ap (assign pos) (*-zeror x) *ℤ-zeror (negsuc x) = ap (assign neg) (*-zeror x) + assign-*l : ∀ {s} x y → assign s (x * y) ≡ assign s x *ℤ pos y + assign-*l {pos} zero y = sym (*ℤ-zerol (pos y)) + assign-*l {pos} (suc x) y = refl + assign-*l {neg} zero y = refl + assign-*l {neg} (suc x) y = refl + + *ℤ-negl : ∀ x y → negℤ x *ℤ y ≡ negℤ (x *ℤ y) + *ℤ-negl posz y = refl + *ℤ-negl (possuc x) y with sign y + ... | pos = neg-assign (suc x * abs y) + ... | neg = neg-assign (suc x * abs y) + *ℤ-negl (negsuc x) y with sign y + ... | pos = neg-assign (suc x * abs y) + ... | neg = neg-assign (suc x * abs y) + private lemma : ∀ x y z → z + y * suc z + x * suc (z + y * suc z) ≡ z + (y + x * suc y) * suc z lemma x y z = nat! @@ -307,14 +416,6 @@ no further comments. dot-is-mul (negsuc zero) (negsuc x) = ap possuc (sym (+-zeror x)) dot-is-mul (negsuc (suc x)) y = sym (*ℤ-negsucl x y ∙ ap (negℤ y +ℤ_) (sym (dot-is-mul (negsuc x) y))) - private - distrib-lemma - : ∀ x y z w → (x +ℤ y) +ℤ (z +ℤ w) ≡ (x +ℤ z) +ℤ (y +ℤ w) - distrib-lemma x y z w = - +ℤ-assoc (x +ℤ y) z w - ·· ap (_+ℤ w) (sym (+ℤ-assoc x y z) ·· ap (x +ℤ_) (+ℤ-commutative y z) ·· +ℤ-assoc x z y) - ·· sym (+ℤ-assoc (x +ℤ z) y w) - dot-distribr : ∀ x y z → dotℤ z (x +ℤ y) ≡ (dotℤ z x) +ℤ (dotℤ z y) dot-distribr x y posz = refl dot-distribr x y (possuc z) = @@ -336,4 +437,35 @@ no further comments. *ℤ-commutative (y +ℤ z) x ·· *ℤ-distribl x y z ·· ap₂ _+ℤ_ (*ℤ-commutative x y) (*ℤ-commutative x z) + + *ℤ-distribr-minus : ∀ x y z → (y -ℤ z) *ℤ x ≡ (y *ℤ x) -ℤ (z *ℤ x) + *ℤ-distribr-minus x y z = *ℤ-distribr x y (negℤ z) ∙ ap (y *ℤ x +ℤ_) (*ℤ-negl z x) + + *ℤ-sucr : ∀ x y → x *ℤ sucℤ y ≡ x *ℤ y +ℤ x + *ℤ-sucr x y = + x *ℤ sucℤ y ≡˘⟨ ap (x *ℤ_) (+ℤ-oner y) ⟩ + x *ℤ (y +ℤ 1) ≡⟨ *ℤ-distribl x y 1 ⟩ + x *ℤ y +ℤ x *ℤ 1 ≡⟨ ap (x *ℤ y +ℤ_) (*ℤ-oner x) ⟩ + x *ℤ y +ℤ x ∎ + + *ℤ-injectiver-possuc : ∀ k x y → x *ℤ possuc k ≡ y *ℤ possuc k → x ≡ y + *ℤ-injectiver-possuc k (pos x) (pos y) p = + ap pos (*-suc-inj k x y (pos-injective (sym (assign-pos _) ·· p ·· assign-pos _))) + *ℤ-injectiver-possuc k (pos x) (negsuc y) p = absurd (pos≠negsuc (sym (assign-pos (x * suc k)) ∙ p)) + *ℤ-injectiver-possuc k (negsuc x) (pos y) p = absurd (negsuc≠pos (p ∙ assign-pos (y * suc k))) + *ℤ-injectiver-possuc k (negsuc x) (negsuc y) p = + ap (assign neg) (*-suc-inj k (suc x) (suc y) (ap suc (negsuc-injective p))) + + *ℤ-injectiver-negsuc : ∀ k x y → x *ℤ negsuc k ≡ y *ℤ negsuc k → x ≡ y + *ℤ-injectiver-negsuc k (pos x) (pos y) p = ap pos (*-suc-inj k x y (pos-injective (negℤ-injective _ _ (sym (assign-neg _) ·· p ·· assign-neg _)))) + *ℤ-injectiver-negsuc k posz (negsuc y) p = absurd (zero≠suc (pos-injective p)) + *ℤ-injectiver-negsuc k (possuc x) (negsuc y) p = absurd (negsuc≠pos p) + *ℤ-injectiver-negsuc k (negsuc x) posz p = absurd (suc≠zero (pos-injective p)) + *ℤ-injectiver-negsuc k (negsuc x) (possuc y) p = absurd (pos≠negsuc p) + *ℤ-injectiver-negsuc k (negsuc x) (negsuc y) p = ap (assign neg) (*-suc-inj k (suc x) (suc y) (ap suc (suc-inj (pos-injective p)))) + + *ℤ-injectiver : ∀ k x y → k ≠ 0 → x *ℤ k ≡ y *ℤ k → x ≡ y + *ℤ-injectiver posz x y k≠0 p = absurd (k≠0 refl) + *ℤ-injectiver (possuc k) x y k≠0 p = *ℤ-injectiver-possuc k x y p + *ℤ-injectiver (negsuc k) x y k≠0 p = *ℤ-injectiver-negsuc k x y p ``` diff --git a/src/Data/Int/Universal.lagda.md b/src/Data/Int/Universal.lagda.md index 101e542a7..c5526e175 100644 --- a/src/Data/Int/Universal.lagda.md +++ b/src/Data/Int/Universal.lagda.md @@ -13,7 +13,7 @@ import Data.Int.HIT as HIT module Data.Int.Universal where ``` -# Universal property of the integers +# Universal property of the integers {defines="universal-property-of-the-integers"} We define and prove a type-theoretic universal property of the integers, which characterises them as the initial pointed set equipped with an diff --git a/src/Data/Nat/DivMod.lagda.md b/src/Data/Nat/DivMod.lagda.md index 6009ce464..8f3e67959 100644 --- a/src/Data/Nat/DivMod.lagda.md +++ b/src/Data/Nat/DivMod.lagda.md @@ -16,7 +16,7 @@ open import Data.Sum.Base module Data.Nat.DivMod where ``` -# Natural division +# Natural division {defines="natural-division"} This module implements the basics of the theory of **division** (not [divisibility], see there) for the natural numbers. In particular, we @@ -160,6 +160,8 @@ d1 /ₙ suc d2 = div-helper 0 d2 d1 d2 _%_ : (d1 d2 : Nat) .⦃ _ : Positive d2 ⦄ → Nat d1 % suc d2 = mod-helper 0 d2 d1 d2 +infixl 9 _/ₙ_ _%_ + abstract private div-mod-lemma : ∀ am ad d n → am + ad * suc (am + n) + d ≡ mod-helper am (am + n) d n + div-helper ad (am + n) d n * suc (am + n) diff --git a/src/Data/Nat/Divisible.lagda.md b/src/Data/Nat/Divisible.lagda.md index b6fefa833..266ae70b5 100644 --- a/src/Data/Nat/Divisible.lagda.md +++ b/src/Data/Nat/Divisible.lagda.md @@ -14,7 +14,7 @@ open import Data.Sum.Base module Data.Nat.Divisible where ``` -# Divisibility {defines="divisibility"} +# Divisibility of natural numbers {defines="divisibility divisible"} In the natural numbers, **divisibility**[^divide] is the property expressing that a given number can be expressed as a multiple of @@ -176,7 +176,7 @@ The only number that divides 1 is 1 itself: ∣-1 {suc (suc n)} (k , p) = *-is-oner k (2 + n) p ``` -## Even and odd numbers +## Even and odd natural numbers A number is **even** if it is divisible by 2, and **odd** otherwise. Note that a number is odd if and only if its successor is even; we take this diff --git a/src/Data/Nat/Order.lagda.md b/src/Data/Nat/Order.lagda.md index 5d4186564..108334fc3 100644 --- a/src/Data/Nat/Order.lagda.md +++ b/src/Data/Nat/Order.lagda.md @@ -235,7 +235,7 @@ that here (it is a more general fact about from {suc x} {suc y} p = s≤s (from (suc-inj p)) ``` -## Well-ordering +## Well-ordering {defines="N-is-well-ordered"} In classical mathematics, the well-ordering principle states that every nonempty subset of the natural numbers has a minimal element. In @@ -256,14 +256,14 @@ implicitly appealing to path induction to reduce this to the case where $p, q : P(n)$] automatically have that $p = q$. ```agda -module _ {ℓ} {P : Nat → Prop ℓ} where - private - minimal-solution : ∀ {ℓ} (P : Nat → Type ℓ) → Type _ - minimal-solution P = Σ[ n ∈ Nat ] (P n × (∀ k → P k → n ≤ k)) - - minimal-solution-unique : is-prop (minimal-solution λ x → ∣ P x ∣) - minimal-solution-unique (n , pn , n-min) (k , pk , k-min) = - Σ-prop-path! (≤-antisym (n-min _ pk) (k-min _ pn)) +minimal-solution : ∀ {ℓ} (P : Nat → Type ℓ) → Type _ +minimal-solution P = Σ[ n ∈ Nat ] (P n × (∀ k → P k → n ≤ k)) + +minimal-solution-unique + : ∀ {ℓ} {P : Nat → Prop ℓ} + → is-prop (minimal-solution λ x → ∣ P x ∣) +minimal-solution-unique (n , pn , n-min) (k , pk , k-min) = + Σ-prop-path! (≤-antisym (n-min _ pk) (k-min _ pn)) ``` The step of the code that actually finds a minimal solution does not @@ -271,30 +271,32 @@ need $P$ to be a predicate: we only need that for actually searching for an inhabited subset. ```agda - min-suc - : ∀ {P : Nat → Type ℓ} → ∀ n → ¬ P 0 - → (∀ k → P (suc k) → n ≤ k) - → ∀ k → P k → suc n ≤ k - min-suc n ¬p0 nmins zero pk = absurd (¬p0 pk) - min-suc zero ¬p0 nmins (suc k) psk = s≤s 0≤x - min-suc (suc n) ¬p0 nmins (suc k) psk = s≤s (nmins k psk) - - ℕ-minimal-solution - : ∀ (P : Nat → Type ℓ) - → (∀ n → Dec (P n)) - → (n : Nat) → P n - → minimal-solution P - ℕ-minimal-solution P decp zero p = 0 , p , λ k _ → 0≤x - ℕ-minimal-solution P decp (suc n) p = case decp zero of λ where - (yes p0) → 0 , p0 , λ k _ → 0≤x - (no ¬p0) → - let (a , pa , x) = ℕ-minimal-solution (P ∘ suc) (decp ∘ suc) n p - in suc a , pa , min-suc {P} a ¬p0 x - - ℕ-well-ordered - : (∀ n → Dec ∣ P n ∣) - → ∃[ n ∈ Nat ] ∣ P n ∣ - → Σ[ n ∈ Nat ] (∣ P n ∣ × (∀ k → ∣ P k ∣ → n ≤ k)) - ℕ-well-ordered P-dec wit = ∥-∥-rec minimal-solution-unique - (λ { (n , p) → ℕ-minimal-solution _ P-dec n p }) wit +min-suc + : ∀ {ℓ} {P : Nat → Type ℓ} + → ∀ n → ¬ P 0 + → (∀ k → P (suc k) → n ≤ k) + → ∀ k → P k → suc n ≤ k +min-suc n ¬p0 nmins zero pk = absurd (¬p0 pk) +min-suc zero ¬p0 nmins (suc k) psk = s≤s 0≤x +min-suc (suc n) ¬p0 nmins (suc k) psk = s≤s (nmins k psk) + +ℕ-minimal-solution + : ∀ {ℓ} (P : Nat → Type ℓ) + → (∀ n → Dec (P n)) + → (n : Nat) → P n + → minimal-solution P +ℕ-minimal-solution P decp zero p = 0 , p , λ k _ → 0≤x +ℕ-minimal-solution P decp (suc n) p = case decp zero of λ where + (yes p0) → 0 , p0 , λ k _ → 0≤x + (no ¬p0) → + let (a , pa , x) = ℕ-minimal-solution (P ∘ suc) (decp ∘ suc) n p + in suc a , pa , min-suc {P = P} a ¬p0 x + +ℕ-well-ordered + : ∀ {ℓ} {P : Nat → Prop ℓ} + → (∀ n → Dec ∣ P n ∣) + → ∃[ n ∈ Nat ] ∣ P n ∣ + → Σ[ n ∈ Nat ] (∣ P n ∣ × (∀ k → ∣ P k ∣ → n ≤ k)) +ℕ-well-ordered {P = P} P-dec wit = ∥-∥-rec (minimal-solution-unique {P = P}) + (λ { (n , p) → ℕ-minimal-solution _ P-dec n p }) wit ``` diff --git a/src/Data/Nat/Properties.lagda.md b/src/Data/Nat/Properties.lagda.md index 71a2563d0..a4b941330 100644 --- a/src/Data/Nat/Properties.lagda.md +++ b/src/Data/Nat/Properties.lagda.md @@ -198,6 +198,11 @@ arithmetic operators: +-≤r x zero = 0≤x +-≤r x (suc y) = subst (λ p → suc y ≤ p) (sym (+-sucr x y)) (s≤s (+-≤r x y)) +monus-≤ : (x y : Nat) → x - y ≤ x +monus-≤ x zero = x≤x +monus-≤ zero (suc y) = 0≤x +monus-≤ (suc x) (suc y) = ≤-sucr (monus-≤ x y) + +-preserves-≤l : (x y z : Nat) → x ≤ y → (z + x) ≤ (z + y) +-preserves-≤l .0 y zero 0≤x = 0≤x +-preserves-≤l .0 y (suc z) 0≤x = diff --git a/src/Data/Set/Coequaliser.lagda.md b/src/Data/Set/Coequaliser.lagda.md index 2e20a318f..c50a763da 100644 --- a/src/Data/Set/Coequaliser.lagda.md +++ b/src/Data/Set/Coequaliser.lagda.md @@ -132,6 +132,10 @@ instance elim! (happly (sf .idsᵉ .to-path h)) Extensional-coeq-map ⦃ sf ⦄ .idsᵉ .to-path-over p = is-prop→pathp (λ i → Pathᵉ-is-hlevel 1 sf (hlevel 2)) _ _ + + Number-Coeq : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → ⦃ Number B ⦄ → {f g : A → B} → Number (Coeq f g) + Number-Coeq {ℓ = ℓ} ⦃ b ⦄ .Number.Constraint n = Lift ℓ (b .Number.Constraint n) + Number-Coeq ⦃ b ⦄ .Number.fromNat n ⦃ lift c ⦄ = inc (b .Number.fromNat n ⦃ c ⦄) ``` --> diff --git a/src/Homotopy/Connectedness.lagda.md b/src/Homotopy/Connectedness.lagda.md index 14223e02a..d2da6921a 100644 --- a/src/Homotopy/Connectedness.lagda.md +++ b/src/Homotopy/Connectedness.lagda.md @@ -18,6 +18,7 @@ module Homotopy.Connectedness where private variable ℓ ℓ' ℓ'' : Level A B : Type ℓ + P : A → Type ℓ x y : A ``` --> @@ -214,9 +215,18 @@ retract→is-n-connected (suc (suc n)) f g h a-conn = (is-n-connected-Tr (suc n) a-conn) ``` + + Since the truncation operator $\| - \|_n$ also preserves products, a remarkably similar argument shows that if $A$ and $B$ are $n$-connected, -then so is $A \times B$ is. +then so is $A \times B$. ```agda ×-is-n-connected @@ -226,6 +236,20 @@ then so is $A \times B$ is. n-Tr-product (×-is-hlevel 0 (n-connected.to n a-conn) (n-connected.to n b-conn)) ``` +We can generalise this to $\Sigma$-types: if $A$ is $n$-connected and +$B$ is an $A$-indexed family of $n$-connected types, then we have +$\|\Sigma_{a : A} B(a)\|_n \simeq \|\Sigma_{a : A} \|B(a)\|_n\|_n +\simeq \|A\|_n \simeq 1$. + +```agda +Σ-is-n-connected + : ∀ n → is-n-connected A n → (∀ a → is-n-connected (P a) n) → is-n-connected (Σ A P) n +Σ-is-n-connected 0 = _ +Σ-is-n-connected (suc n) a-conn b-conn = n-connected.from n $ Equiv→is-hlevel 0 + (n-Tr-Σ ∙e n-Tr-≃ (Σ-contract λ _ → n-connected.to n (b-conn _))) + (n-connected.to n a-conn) +``` + Finally, we show the dual of two properties of truncations: if $A$ is $(1+n)$-connected, then each path space $x \equiv_A y$ is $n$-connected; And if $A$ is $(2+n)$-connected, then it is also $(1+n$) connected. This diff --git a/src/Homotopy/Connectedness/Automation.lagda.md b/src/Homotopy/Connectedness/Automation.lagda.md index 569b7959a..a0c2ecdc3 100644 --- a/src/Homotopy/Connectedness/Automation.lagda.md +++ b/src/Homotopy/Connectedness/Automation.lagda.md @@ -64,12 +64,12 @@ instance Connected-⊤ {n} = conn-instance (n-connected-∥-∥.from n (is-contr→is-connected (hlevel 0))) - Connected-× - : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {n} - → ⦃ Connected A n ⦄ → ⦃ Connected B n ⦄ - → Connected (A × B) n - Connected-× {n = n} ⦃ conn-instance ac ⦄ ⦃ conn-instance bc ⦄ = conn-instance - (×-is-n-connected n ac bc) + Connected-Σ + : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {n} + → ⦃ Connected A n ⦄ → ⦃ ∀ {a} → Connected (B a) n ⦄ + → Connected (Σ A B) n + Connected-Σ {n = n} ⦃ conn-instance ac ⦄ ⦃ bc ⦄ = conn-instance + (Σ-is-n-connected n ac λ _ → Connected.has-n-connected bc) Connected-Path : ∀ {ℓ} {A : Type ℓ} {x y : A} {n} @@ -77,4 +77,11 @@ instance → Connected (Path A x y) n Connected-Path {n = n} ⦃ conn-instance ac ⦄ = conn-instance (Path-is-connected n ac) + + Connected-Lift + : ∀ {ℓ ℓ'} {A : Type ℓ} {n} + → ⦃ Connected A n ⦄ + → Connected (Lift ℓ' A) n + Connected-Lift {n = n} ⦃ conn-instance ac ⦄ = conn-instance + (is-n-connected-≃ n (Lift-≃ e⁻¹) ac) ``` diff --git a/src/Homotopy/Space/Circle.lagda.md b/src/Homotopy/Space/Circle.lagda.md index b8d22cdc6..ac74cc9a7 100644 --- a/src/Homotopy/Space/Circle.lagda.md +++ b/src/Homotopy/Space/Circle.lagda.md @@ -4,9 +4,9 @@ open import 1Lab.Path.Reasoning open import 1Lab.Prelude open import Algebra.Group.Cat.FinitelyComplete +open import Algebra.Group.Instances.Integers open import Algebra.Group.Cat.Base open import Algebra.Group.Homotopy -open import Algebra.Group.Ab open import Algebra.Group open import Data.Set.Truncation @@ -299,10 +299,10 @@ loopⁿ-+ a = Integers.induction Int-integers loopⁿ (a +ℤ sucℤ b) ≡ loopⁿ a ∙ loopⁿ (sucℤ b) ≃∎ π₁S¹≡ℤ : π₁Groupoid.π₁ S¹∙ S¹-is-groupoid ≡ ℤ -π₁S¹≡ℤ = sym $ ∫-Path Groups-equational - (total-hom (Equiv.from ΩS¹≃integers ∘ Lift.lower) - (record { pres-⋆ = λ (lift a) (lift b) → loopⁿ-+ a b })) - (∙-is-equiv (Lift-≃ .snd) ((ΩS¹≃integers e⁻¹) .snd)) +π₁S¹≡ℤ = sym $ ∫-Path + (total-hom (Equiv.from ΩS¹≃integers) + (record { pres-⋆ = loopⁿ-+ })) + ((ΩS¹≃integers e⁻¹) .snd) ``` Furthermore, since the loop space of the circle is a set, we automatically @@ -314,7 +314,7 @@ get that all of its higher homotopy groups are trivial. Ωⁿ⁺²S¹-is-contr (suc n) = Path-is-hlevel 0 (Ωⁿ⁺²S¹-is-contr n) πₙ₊₂S¹≡0 : ∀ n → πₙ₊₁ (suc n) S¹∙ ≡ Zero-group {lzero} -πₙ₊₂S¹≡0 n = ∫-Path Groups-equational +πₙ₊₂S¹≡0 n = ∫-Path (Zero-group-is-terminal _ .centre) (is-contr→≃ (is-contr→∥-∥₀-is-contr (Ωⁿ⁺²S¹-is-contr n)) (hlevel 0) .snd) ``` diff --git a/src/Homotopy/Space/Delooping.lagda.md b/src/Homotopy/Space/Delooping.lagda.md index 6a7314322..fc218d3e1 100644 --- a/src/Homotopy/Space/Delooping.lagda.md +++ b/src/Homotopy/Space/Delooping.lagda.md @@ -284,7 +284,7 @@ group of `Deloop`{.Agda} is `G`, which is what we wanted. G≃ΩB = Iso→Equiv (decode base , iso (encode base) encode→decode (decode→encode base)) G≡π₁B : G ≡ πₙ₊₁ 0 (Deloop , base) - G≡π₁B = ∫-Path Groups-equational + G≡π₁B = ∫-Path (total-hom (λ x → inc (path x)) record { pres-⋆ = λ x y → ap ∥_∥₀.inc (path-∙ _ _) }) (∙-is-equiv (G≃ΩB .snd) (∥-∥₀-idempotent (squash base base))) diff --git a/src/Homotopy/Truncation.lagda.md b/src/Homotopy/Truncation.lagda.md index 3b15b1e74..60deeb849 100644 --- a/src/Homotopy/Truncation.lagda.md +++ b/src/Homotopy/Truncation.lagda.md @@ -183,6 +183,11 @@ n-Tr-rec! → ⦃ hl : H-Level B (suc n) ⦄ → (A → B) → n-Tr A (suc n) → B n-Tr-rec! = n-Tr-elim (λ _ → _) (λ _ → hlevel _) + +n-Tr-map + : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {n} + → (A → B) → n-Tr A (suc n) → n-Tr B (suc n) +n-Tr-map f = n-Tr-rec! (inc ∘ f) ``` --> @@ -267,7 +272,7 @@ n-Tr-univ n b-hl = Iso→Equiv λ where n-Tr-product : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {n} → n-Tr (A × B) (suc n) ≃ (n-Tr A (suc n) × n-Tr B (suc n)) -n-Tr-product {A = A} {B} {n} = distrib , distrib-is-equiv module n-Tr-product where +n-Tr-product {A = A} {B} {n} = distrib , distrib-is-equiv where distrib : n-Tr (A × B) (suc n) → n-Tr A (suc n) × n-Tr B (suc n) distrib (inc (x , y)) = inc x , inc y distrib (hub r) .fst = hub λ s → distrib (r s) .fst @@ -291,4 +296,24 @@ n-Tr-product {A = A} {B} {n} = distrib , distrib-is-equiv module n-Tr-product wh distrib-is-iso .linv = n-Tr-elim! _ λ x → refl distrib-is-equiv = is-iso→is-equiv distrib-is-iso + +n-Tr-Σ + : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {n} + → n-Tr (Σ A B) (suc n) ≃ n-Tr (Σ A λ a → n-Tr (B a) (suc n)) (suc n) +n-Tr-Σ {A = A} {B} {n} = Iso→Equiv is where + is : Iso _ _ + is .fst = n-Tr-map (Σ-map id inc) + is .snd .is-iso.inv = n-Tr-rec! λ (a , b) → n-Tr-map (a ,_) b + is .snd .is-iso.rinv = n-Tr-elim! _ λ (a , b) → n-Tr-elim! (λ b → n-Tr-map (Σ-map id inc) (n-Tr-map (a ,_) b) ≡ inc (a , b)) (λ _ → refl) b + is .snd .is-iso.linv = n-Tr-elim! _ λ _ → refl + +n-Tr-≃ + : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {n} + → (e : A ≃ B) → n-Tr A (suc n) ≃ n-Tr B (suc n) +n-Tr-≃ e = Iso→Equiv is where + is : Iso _ _ + is .fst = n-Tr-map (e .fst) + is .snd .is-iso.inv = n-Tr-map (Equiv.from e) + is .snd .is-iso.rinv = elim! λ _ → ap inc (Equiv.ε e _) + is .snd .is-iso.linv = elim! λ _ → ap inc (Equiv.η e _) ``` From afcf848d367f906d6d07d1612fbd41d7dd8c978e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Am=C3=A9lia?= Date: Sun, 8 Sep 2024 13:41:17 -0300 Subject: [PATCH 6/9] ring localisations and rational numbers (#429) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit numbers 👎 --- src/1Lab/Equiv.lagda.md | 4 +- src/1Lab/HLevel/Closure.lagda.md | 4 +- src/1Lab/HLevel/Universe.lagda.md | 2 +- src/1Lab/Reflection/Deriving/Show.agda | 1 + src/1Lab/Type.lagda.md | 2 + src/1Lab/Underlying.agda | 2 +- src/1Lab/Univalence.lagda.md | 6 +- src/Algebra/Group.lagda.md | 10 +- src/Algebra/Group/Ab/Sum.lagda.md | 2 +- src/Algebra/Group/Action.lagda.md | 5 +- src/Algebra/Group/Cat/Base.lagda.md | 3 +- src/Algebra/Group/Cat/Monadic.lagda.md | 3 +- .../Group/Concrete/Semidirect.lagda.md | 4 +- src/Algebra/Group/Free/Product.lagda.md | 5 +- src/Algebra/Group/Instances/Cyclic.lagda.md | 7 +- src/Algebra/Group/Instances/Dihedral.lagda.md | 4 +- src/Algebra/Group/Instances/Integers.lagda.md | 4 +- src/Algebra/Group/Semidirect.lagda.md | 3 +- src/Algebra/Group/Subgroup.lagda.md | 2 +- src/Algebra/Magma.lagda.md | 2 +- .../Magma/Unital/EckmannHilton.lagda.md | 2 + src/Algebra/Monoid.lagda.md | 2 +- src/Algebra/Monoid/Category.lagda.md | 1 + src/Algebra/Monoid/Reasoning.lagda.md | 86 +++ src/Algebra/Prelude.agda | 31 - src/Algebra/Ring.lagda.md | 37 +- src/Algebra/Ring/Cat/Initial.lagda.md | 183 ++--- src/Algebra/Ring/Commutative.lagda.md | 21 +- src/Algebra/Ring/Ideal.lagda.md | 26 +- src/Algebra/Ring/Localisation.lagda.md | 431 ++++++++++++ src/Algebra/Ring/Module/Action.lagda.md | 16 + src/Algebra/Ring/Module/Free.lagda.md | 4 +- src/Algebra/Ring/Module/Multilinear.lagda.md | 1 - src/Algebra/Ring/Quotient.lagda.md | 13 +- src/Algebra/Ring/Reasoning.lagda.md | 51 ++ src/Algebra/Ring/Solver.agda | 82 ++- src/Cat/Abelian/Base.lagda.md | 24 +- src/Cat/Abelian/Instances/Functor.lagda.md | 4 +- src/Cat/CartesianClosed/Instances/PSh.agda | 6 +- src/Cat/Diagram/Coproduct/Indexed.lagda.md | 2 +- src/Cat/Diagram/Everything.agda | 30 - src/Cat/Diagram/Product/Indexed.lagda.md | 2 +- src/Cat/Displayed/Cartesian/Indexing.lagda.md | 1 - src/Cat/Displayed/Instances/Opposite.lagda.md | 1 - src/Cat/Instances/OFE/Coproduct.lagda.md | 20 +- src/Cat/Instances/OFE/Product.lagda.md | 4 +- .../Sets/Counterexamples/SelfDual.lagda.md | 2 +- src/Cat/Monoidal/Diagram/Monoid.lagda.md | 10 +- .../Diagram/Monoid/Representable.lagda.md | 4 +- src/Data/Dec/Base.lagda.md | 8 + src/Data/Fin/Closure.lagda.md | 2 +- src/Data/Fin/Product.lagda.md | 51 ++ src/Data/Int/Base.lagda.md | 41 +- src/Data/Int/DivMod.lagda.md | 6 +- src/Data/Int/Order.lagda.md | 44 +- src/Data/Int/Properties.lagda.md | 89 ++- src/Data/List/Properties.lagda.md | 4 +- src/Data/Nat/Base.lagda.md | 4 + src/Data/Nat/Divisible.lagda.md | 3 + src/Data/Nat/Divisible/GCD.lagda.md | 18 + src/Data/Nat/Prime.lagda.md | 12 - src/Data/Nat/Properties.lagda.md | 39 +- src/Data/Rational/Base.lagda.md | 626 ++++++++++++++++++ src/Data/Rational/Order.lagda.md | 239 +++++++ src/Data/Set/Coequaliser.lagda.md | 36 +- src/Data/Set/Coequaliser/Split.lagda.md | 116 ++++ src/Data/Set/Material.lagda.md | 4 +- src/Data/Sum/Properties.lagda.md | 4 +- src/Order/DCPO/Free.lagda.md | 10 +- src/Order/DCPO/Pointed.lagda.md | 4 +- src/Order/Diagram/Glb.lagda.md | 10 +- src/Order/Diagram/Lub.lagda.md | 10 +- src/Order/Instances/Lift.lagda.md | 2 +- src/Order/Semilattice/Join/Reasoning.lagda.md | 2 + src/Order/Semilattice/Meet/Reasoning.lagda.md | 2 + src/Topoi/Base.lagda.md | 14 +- src/Topoi/Classifying/Diaconescu.lagda.md | 5 +- src/Topoi/Reasoning.lagda.md | 5 +- src/preamble.tex | 4 + 79 files changed, 2208 insertions(+), 378 deletions(-) create mode 100644 src/Algebra/Monoid/Reasoning.lagda.md delete mode 100644 src/Algebra/Prelude.agda create mode 100644 src/Algebra/Ring/Localisation.lagda.md create mode 100644 src/Algebra/Ring/Reasoning.lagda.md delete mode 100644 src/Cat/Diagram/Everything.agda create mode 100644 src/Data/Rational/Base.lagda.md create mode 100644 src/Data/Rational/Order.lagda.md create mode 100644 src/Data/Set/Coequaliser/Split.lagda.md diff --git a/src/1Lab/Equiv.lagda.md b/src/1Lab/Equiv.lagda.md index 89bff45ec..783681231 100644 --- a/src/1Lab/Equiv.lagda.md +++ b/src/1Lab/Equiv.lagda.md @@ -945,7 +945,7 @@ Lift-ap Lift-ap (f , eqv) .fst (lift x) = lift (f x) Lift-ap f .snd .is-eqv (lift y) .centre = lift (Equiv.from f y) , ap lift (Equiv.ε f y) Lift-ap f .snd .is-eqv (lift y) .paths (lift x , q) i = lift (p i .fst) , λ j → lift (p i .snd j) - where p = f .snd .is-eqv y .paths (x , ap Lift.lower q) + where p = f .snd .is-eqv y .paths (x , ap lower q) ``` ### Involutions @@ -1090,7 +1090,7 @@ syntax ≃⟨⟩-syntax x q p = x ≃⟨ p ⟩ q lift-inj : ∀ {ℓ ℓ'} {A : Type ℓ} {a b : A} → lift {ℓ = ℓ'} a ≡ lift {ℓ = ℓ'} b → a ≡ b -lift-inj p = ap Lift.lower p +lift-inj p = ap lower p -- Fibres of composites are given by pairs of fibres. fibre-∘-≃ diff --git a/src/1Lab/HLevel/Closure.lagda.md b/src/1Lab/HLevel/Closure.lagda.md index 2c0a90224..d626fec0f 100644 --- a/src/1Lab/HLevel/Closure.lagda.md +++ b/src/1Lab/HLevel/Closure.lagda.md @@ -266,13 +266,13 @@ successor universe: ```agda opaque Lift-is-hlevel : ∀ n → is-hlevel A n → is-hlevel (Lift ℓ' A) n - Lift-is-hlevel n a-hl = retract→is-hlevel n lift Lift.lower (λ _ → refl) a-hl + Lift-is-hlevel n a-hl = retract→is-hlevel n lift lower (λ _ → refl) a-hl ``` diff --git a/src/1Lab/HLevel/Universe.lagda.md b/src/1Lab/HLevel/Universe.lagda.md index a7177e62f..2191c2a42 100644 --- a/src/1Lab/HLevel/Universe.lagda.md +++ b/src/1Lab/HLevel/Universe.lagda.md @@ -193,7 +193,7 @@ Prop ℓ = n-Type ℓ 1 ```agda ¬Set-is-prop : ¬ is-prop (Set ℓ) ¬Set-is-prop prop = - Lift.lower $ + lower $ transport (ap ∣_∣ (prop (el (Lift _ ⊤) (hlevel 2)) (el (Lift _ ⊥) (hlevel 2)))) (lift tt) ``` --> diff --git a/src/1Lab/Reflection/Deriving/Show.agda b/src/1Lab/Reflection/Deriving/Show.agda index 99e9927b7..f3d740351 100644 --- a/src/1Lab/Reflection/Deriving/Show.agda +++ b/src/1Lab/Reflection/Deriving/Show.agda @@ -11,6 +11,7 @@ open import Data.Char.Base open import Data.Fin.Base open import Data.Vec.Base hiding (map) open import Data.Nat.Base +open import Data.Int.Base using (Int ; pos ; negsuc) open import Meta.Foldable open import Meta.Append diff --git a/src/1Lab/Type.lagda.md b/src/1Lab/Type.lagda.md index d0f4e1ad9..8543066e0 100644 --- a/src/1Lab/Type.lagda.md +++ b/src/1Lab/Type.lagda.md @@ -77,6 +77,8 @@ record Lift {a} ℓ (A : Type a) : Type (a ⊔ ℓ) where @@ -93,9 +93,7 @@ give the unit, both on the left and on the right: underlying-monoid = A , record { identity = unit ; _⋆_ = _*_ ; has-is-monoid = has-is-monoid } - open Cat.Reasoning (B (underlying-monoid .snd)) - hiding (id ; assoc ; idl ; idr ; invr ; invl ; to ; from ; inverses ; _∘_) - public + open Mon underlying-monoid public ``` --> diff --git a/src/Algebra/Group/Ab/Sum.lagda.md b/src/Algebra/Group/Ab/Sum.lagda.md index f2b0899cc..2d88deca7 100644 --- a/src/Algebra/Group/Ab/Sum.lagda.md +++ b/src/Algebra/Group/Ab/Sum.lagda.md @@ -3,10 +3,10 @@ open import Algebra.Group.Cat.FinitelyComplete open import Algebra.Group.Cat.Base open import Algebra.Group.Ab -open import Algebra.Prelude open import Algebra.Group open import Cat.Diagram.Product +open import Cat.Prelude ``` --> diff --git a/src/Algebra/Group/Action.lagda.md b/src/Algebra/Group/Action.lagda.md index 62e949ddb..43a66ec19 100644 --- a/src/Algebra/Group/Action.lagda.md +++ b/src/Algebra/Group/Action.lagda.md @@ -3,13 +3,16 @@ open import Algebra.Group.Cat.FinitelyComplete open import Algebra.Group.Cat.Base open import Algebra.Group.Solver -open import Algebra.Prelude open import Algebra.Group open import Cat.Instances.Delooping +open import Cat.Instances.Functor open import Cat.Instances.Sets +open import Cat.Diagram.Zero +open import Cat.Prelude import Cat.Functor.Reasoning as Functor-kit +import Cat.Reasoning as Cat ``` --> diff --git a/src/Algebra/Group/Cat/Base.lagda.md b/src/Algebra/Group/Cat/Base.lagda.md index 2943014ab..722e70d41 100644 --- a/src/Algebra/Group/Cat/Base.lagda.md +++ b/src/Algebra/Group/Cat/Base.lagda.md @@ -1,7 +1,6 @@ diff --git a/src/Algebra/Group/Cat/Monadic.lagda.md b/src/Algebra/Group/Cat/Monadic.lagda.md index 6c0df60e4..fa7af9e1f 100644 --- a/src/Algebra/Group/Cat/Monadic.lagda.md +++ b/src/Algebra/Group/Cat/Monadic.lagda.md @@ -2,7 +2,6 @@ ```agda open import Algebra.Group.Cat.Base open import Algebra.Group.Free -open import Algebra.Prelude open import Algebra.Monoid open import Algebra.Group @@ -10,7 +9,9 @@ open import Cat.Functor.Adjoint.Monadic open import Cat.Functor.Adjoint.Monad open import Cat.Functor.Equivalence open import Cat.Functor.Properties +open import Cat.Functor.Adjoint open import Cat.Diagram.Monad +open import Cat.Prelude import Algebra.Group.Cat.Base as Grp diff --git a/src/Algebra/Group/Concrete/Semidirect.lagda.md b/src/Algebra/Group/Concrete/Semidirect.lagda.md index c0d3e7787..d1545c260 100644 --- a/src/Algebra/Group/Concrete/Semidirect.lagda.md +++ b/src/Algebra/Group/Concrete/Semidirect.lagda.md @@ -4,9 +4,11 @@ open import Algebra.Group.Semidirect open import Algebra.Group.Cat.Base open import Algebra.Group.Concrete open import Algebra.Group.Action -open import Algebra.Prelude open import Algebra.Group +open import Cat.Univalent +open import Cat.Prelude + open import Homotopy.Connectedness open ConcreteGroup diff --git a/src/Algebra/Group/Free/Product.lagda.md b/src/Algebra/Group/Free/Product.lagda.md index f31e942f3..3075bc605 100644 --- a/src/Algebra/Group/Free/Product.lagda.md +++ b/src/Algebra/Group/Free/Product.lagda.md @@ -4,10 +4,13 @@ open import 1Lab.Reflection.Induction open import Algebra.Group.Cat.FinitelyComplete open import Algebra.Group.Cat.Base -open import Algebra.Prelude open import Algebra.Group open import Cat.Diagram.Colimit.Finite +open import Cat.Diagram.Coproduct +open import Cat.Diagram.Pushout +open import Cat.Diagram.Zero +open import Cat.Prelude open Finitely-cocomplete open is-group-hom diff --git a/src/Algebra/Group/Instances/Cyclic.lagda.md b/src/Algebra/Group/Instances/Cyclic.lagda.md index e45f16113..ec0aed61c 100644 --- a/src/Algebra/Group/Instances/Cyclic.lagda.md +++ b/src/Algebra/Group/Instances/Cyclic.lagda.md @@ -5,15 +5,16 @@ open import Algebra.Group.Instances.Integers open import Algebra.Group.Cat.Base open import Algebra.Group.Subgroup open import Algebra.Group.Ab -open import Algebra.Prelude open import Algebra.Group +open import Cat.Prelude + open import Data.Int.Divisible open import Data.Int.Universal open import Data.Fin.Closure open import Data.Int.DivMod open import Data.Fin -open import Data.Int +open import Data.Int hiding (Positive) open import Data.Nat open represents-subgroup @@ -82,7 +83,7 @@ _·ℤ : ∀ (n : Nat) → normal-subgroup ℤ λ i → el (n ∣ℤ i) (∣ℤ- where x≡y+x-y : x ≡ y +ℤ (x -ℤ y) x≡y+x-y = - x ≡⟨ ℤ.insertl {h = y} (ℤ.inverser {x = y}) ⟩ + x ≡⟨ ℤ.insertl {y} (ℤ.inverser {x = y}) ⟩ y +ℤ (negℤ y +ℤ x) ≡⟨ ap (y +ℤ_) (+ℤ-commutative (negℤ y) x) ⟩ y +ℤ (x -ℤ y) ∎ diff --git a/src/Algebra/Group/Instances/Dihedral.lagda.md b/src/Algebra/Group/Instances/Dihedral.lagda.md index 436b5e0ee..ce6ae3c1b 100644 --- a/src/Algebra/Group/Instances/Dihedral.lagda.md +++ b/src/Algebra/Group/Instances/Dihedral.lagda.md @@ -6,7 +6,9 @@ open import Algebra.Group.Semidirect open import Algebra.Group.Cat.Base open import Algebra.Group.Action open import Algebra.Group.Ab -open import Algebra.Prelude + +open import Cat.Functor.Base +open import Cat.Prelude ``` --> diff --git a/src/Algebra/Group/Instances/Integers.lagda.md b/src/Algebra/Group/Instances/Integers.lagda.md index 50718c63e..23df0c39e 100644 --- a/src/Algebra/Group/Instances/Integers.lagda.md +++ b/src/Algebra/Group/Instances/Integers.lagda.md @@ -2,14 +2,14 @@ ```agda open import Algebra.Group.Cat.Base open import Algebra.Group.Ab -open import Algebra.Prelude open import Algebra.Group open import Cat.Functor.Adjoint +open import Cat.Prelude open import Data.Int.Universal open import Data.Nat.Order -open import Data.Int +open import Data.Int hiding (Positive) open import Data.Nat open is-group-hom diff --git a/src/Algebra/Group/Semidirect.lagda.md b/src/Algebra/Group/Semidirect.lagda.md index 97e1c7e2d..c79ae3b28 100644 --- a/src/Algebra/Group/Semidirect.lagda.md +++ b/src/Algebra/Group/Semidirect.lagda.md @@ -3,9 +3,10 @@ open import Algebra.Group.Cat.FinitelyComplete open import Algebra.Group.Cat.Base open import Algebra.Group.Action -open import Algebra.Prelude open import Algebra.Group +open import Cat.Prelude + open is-group-hom open make-group ``` diff --git a/src/Algebra/Group/Subgroup.lagda.md b/src/Algebra/Group/Subgroup.lagda.md index 06d14ac30..964e536b2 100644 --- a/src/Algebra/Group/Subgroup.lagda.md +++ b/src/Algebra/Group/Subgroup.lagda.md @@ -2,13 +2,13 @@ ```agda open import Algebra.Group.Cat.FinitelyComplete open import Algebra.Group.Cat.Base -open import Algebra.Prelude open import Algebra.Group open import Cat.Diagram.Equaliser.Kernel open import Cat.Diagram.Coequaliser open import Cat.Diagram.Equaliser open import Cat.Diagram.Zero +open import Cat.Prelude open import Data.Power diff --git a/src/Algebra/Magma.lagda.md b/src/Algebra/Magma.lagda.md index c7a0e0cf3..21e49c610 100644 --- a/src/Algebra/Magma.lagda.md +++ b/src/Algebra/Magma.lagda.md @@ -64,7 +64,7 @@ binary operation `⋆`, on which no further laws are imposed. underlying-set : Set ℓ underlying-set = el _ has-is-set - instance + opaque instance magma-hlevel : ∀ {n} → H-Level A (2 + n) magma-hlevel = basic-instance 2 has-is-set diff --git a/src/Algebra/Magma/Unital/EckmannHilton.lagda.md b/src/Algebra/Magma/Unital/EckmannHilton.lagda.md index 84d502ebc..5140eee04 100644 --- a/src/Algebra/Magma/Unital/EckmannHilton.lagda.md +++ b/src/Algebra/Magma/Unital/EckmannHilton.lagda.md @@ -5,6 +5,8 @@ open import 1Lab.Prelude open import Algebra.Magma.Unital open import Algebra.Semigroup open import Algebra.Monoid + +open is-monoid ``` --> diff --git a/src/Algebra/Monoid.lagda.md b/src/Algebra/Monoid.lagda.md index fc71bf790..7b0fefe8f 100644 --- a/src/Algebra/Monoid.lagda.md +++ b/src/Algebra/Monoid.lagda.md @@ -46,7 +46,7 @@ record is-monoid (id : A) (_⋆_ : A → A → A) : Type (level-of A) where idl : {x : A} → id ⋆ x ≡ x idr : {x : A} → x ⋆ id ≡ x -open is-monoid public +open is-monoid ``` The condition of $(A, 0, \star)$ defining a monoid is a proposition, so diff --git a/src/Algebra/Monoid/Category.lagda.md b/src/Algebra/Monoid/Category.lagda.md index 20958c37d..f1b16ca9b 100644 --- a/src/Algebra/Monoid/Category.lagda.md +++ b/src/Algebra/Monoid/Category.lagda.md @@ -25,6 +25,7 @@ module Algebra.Monoid.Category where ``` open Precategory open is-semigroup +open is-monoid open is-magma open Monoid-hom open Monoid-on diff --git a/src/Algebra/Monoid/Reasoning.lagda.md b/src/Algebra/Monoid/Reasoning.lagda.md new file mode 100644 index 000000000..d4ba9b81e --- /dev/null +++ b/src/Algebra/Monoid/Reasoning.lagda.md @@ -0,0 +1,86 @@ + + +```agda +module Algebra.Monoid.Reasoning {ℓ} (M : Monoid ℓ) where +``` + +# Reasoning combinators for monoids + + + +```agda +id-comm : a * 1m ≡ 1m * a +id-comm = idr ∙ sym idl + +id-comm-sym : 1m * a ≡ a * 1m +id-comm-sym = idl ∙ sym idr + +module _ (p : x ≡ 1m) where + eliml : x * a ≡ a + eliml = ap₂ _*_ p refl ∙ idl + + elimr : a * x ≡ a + elimr = ap₂ _*_ refl p ∙ idr + + introl : a ≡ x * a + introl = sym eliml + + intror : a ≡ a * x + intror = sym elimr + +module _ (p : x * y ≡ z) where + pulll : x * (y * a) ≡ z * a + pulll = associative ∙ ap₂ _*_ p refl + + pullr : (a * x) * y ≡ a * z + pullr = sym associative ∙ ap₂ _*_ refl p + +module _ (p : z ≡ x * y) where + pushl : z * a ≡ x * (y * a) + pushl = sym (pulll (sym p)) + + pushr : a * z ≡ (a * x) * y + pushr = sym (pullr (sym p)) + +module _ (p : w * x ≡ y * z) where + extendl : w * (x * a) ≡ y * (z * a) + extendl = pulll refl ∙ ap₂ _*_ p refl ∙ pullr refl + + extendr : (a * w) * x ≡ (a * y) * z + extendr = pullr refl ∙ ap₂ _*_ refl p ∙ pulll refl + +module _ (p : x * y ≡ 1m) where + cancell : x * (y * a) ≡ a + cancell = pulll p ∙ idl + + cancelr : (a * x) * y ≡ a + cancelr = pullr p ∙ idr + + insertl : a ≡ x * (y * a) + insertl = sym cancell + + insertr : a ≡ (a * x) * y + insertr = sym cancelr + +lswizzle : g ≡ h * i → f * h ≡ 1m → f * g ≡ i +lswizzle p q = ap₂ _*_ refl p ∙ cancell q + +rswizzle : g ≡ i * h → h * f ≡ 1m → g * f ≡ i +rswizzle p q = ap₂ _*_ p refl ∙ cancelr q + +swizzle : f * g ≡ h * i → g * g' ≡ 1m → h' * h ≡ 1m → h' * f ≡ i * g' +swizzle p q r = lswizzle (sym (associative ∙ rswizzle (sym p) q)) r +``` diff --git a/src/Algebra/Prelude.agda b/src/Algebra/Prelude.agda deleted file mode 100644 index f43e9435d..000000000 --- a/src/Algebra/Prelude.agda +++ /dev/null @@ -1,31 +0,0 @@ -module Algebra.Prelude where - -open import Cat.Instances.Shape.Terminal public -open import Cat.Functor.FullSubcategory public -open import Cat.Instances.Product public -open import Cat.Instances.Functor public -open import Cat.Instances.Comma public -open import Cat.Instances.Slice public -open import Cat.Functor.Adjoint public -open import Cat.Functor.Base public -open import Cat.Functor.Hom public -open import Cat.Univalent public -open import Cat.Prelude hiding (_+_ ; _-_ ; _*_) public - -open import Cat.Diagram.Coequaliser public -open import Cat.Diagram.Coproduct public -open import Cat.Diagram.Equaliser public -open import Cat.Diagram.Pullback public -open import Cat.Diagram.Terminal public -open import Cat.Diagram.Initial public -open import Cat.Diagram.Pushout public -open import Cat.Diagram.Product public -open import Cat.Diagram.Image public -open import Cat.Diagram.Zero public - -import Cat.Reasoning - -open import Cat.Displayed.Univalence.Thin public - -module Cat {o ℓ} (C : Precategory o ℓ) where - open Cat.Reasoning C public diff --git a/src/Algebra/Ring.lagda.md b/src/Algebra/Ring.lagda.md index d3ebbfe8d..167777f48 100644 --- a/src/Algebra/Ring.lagda.md +++ b/src/Algebra/Ring.lagda.md @@ -1,18 +1,20 @@ @@ -72,7 +74,7 @@ record is-ring {ℓ} {R : Type ℓ} (1r : R) (_*_ _+_ : R → R → R) : Type ) public - additive-group : Group ℓ + additive-group : Σ (Set ℓ) (λ x → Group-on ⌞ x ⌟) ∣ additive-group .fst ∣ = R additive-group .fst .is-tr = is-abelian-group.has-is-set +-group additive-group .snd .Group-on._⋆_ = _+_ @@ -84,27 +86,13 @@ record is-ring {ℓ} {R : Type ℓ} (1r : R) (_*_ _+_ : R → R → R) : Type group .snd .Abelian-group-on._*_ = _+_ group .snd .Abelian-group-on.has-is-ab = +-group - Ringoid : Ab-category (B record { _⋆_ = _*_ ; has-is-monoid = *-monoid }) - Ringoid .Ab-category.Abelian-group-on-hom _ _ = record { has-is-ab = +-group } - Ringoid .Ab-category.∘-linear-l f g h = sym *-distribr - Ringoid .Ab-category.∘-linear-r f g h = sym *-distribl - - private - module ringoid = Ab-category Ringoid - using ( ∘-zero-l ; ∘-zero-r ; neg-∘-l ; neg-∘-r ; ∘-minus-l ; ∘-minus-r ) - - open ringoid renaming - ( ∘-zero-l to *-zerol - ; ∘-zero-r to *-zeror - ; neg-∘-l to neg-*-l - ; neg-∘-r to neg-*-r - ; ∘-minus-l to *-minus-l - ; ∘-minus-r to *-minus-r - ) - public + multiplicative-monoid : Monoid ℓ + multiplicative-monoid .fst = R + multiplicative-monoid .snd = record { has-is-monoid = *-monoid } - module m = Cat.Reasoning (B record { _⋆_ = _*_ ; has-is-monoid = *-monoid }) + module m = Mon multiplicative-monoid module a = Abelian-group-on record { has-is-ab = +-group } + hiding (_*_ ; 1g ; _⁻¹) record Ring-on {ℓ} (R : Type ℓ) : Type ℓ where field @@ -246,6 +234,7 @@ record make-ring {ℓ} (R : Type ℓ) : Type ℓ where to-ring-on : Ring-on R to-ring-on = ring where open is-ring hiding (-_ ; +-invr ; +-invl ; *-distribl ; *-distribr ; *-idl ; *-idr ; +-idl ; +-idr) + open is-monoid -- All in copatterns to prevent the unfolding from exploding on you ring : Ring-on R diff --git a/src/Algebra/Ring/Cat/Initial.lagda.md b/src/Algebra/Ring/Cat/Initial.lagda.md index 0e77e0bf6..cf53b4bd0 100644 --- a/src/Algebra/Ring/Cat/Initial.lagda.md +++ b/src/Algebra/Ring/Cat/Initial.lagda.md @@ -1,14 +1,17 @@ @@ -69,13 +72,13 @@ prove...], so here it is: ```agda Int-is-initial : is-initial (Rings ℓ) Liftℤ Int-is-initial R = contr z→r λ x → ext (lemma x) where - module R = Ring-on (R .snd) + module R = Kit R ``` Note that we treat 1 with care: we could have this map 1 to `1r + 0r`, -but this results in worse definitional behaviour when actually using the embedding. -This will result in a bit more work right now, but is work worth doing. - +but this results in worse definitional behaviour when actually using the +embedding. This will result in a bit more work right now, but is work +worth doing. ```agda e : Nat → ⌞ R ⌟ @@ -142,78 +145,60 @@ this is a ring homomorphism... which involves a mountain of annoying algebra, so I won't comment on it too much: it can be worked out on paper, following the ring laws. -Note that we special case `diff x 0` here for better definitional -behaviour of the embedding. - ```agda - ℤ↪R : Int → ⌞ R ⌟ - ℤ↪R (diff x zero) = e x - ℤ↪R (diff x (suc y)) = e x R.- e (suc y) - ℤ↪R (Int.quot m zero i) = along i $ - e m ≡⟨ R.intror R.+-invr ⟩ - e m R.+ (R.1r R.- R.1r) ≡⟨ R.+-associative ⟩ - (e m R.+ R.1r) R.- R.1r ≡⟨ ap (R._- R.1r) R.+-commutes ⟩ - (R.1r R.+ e m) R.- R.1r ≡˘⟨ ap (R._- R.1r) (e-suc m) ⟩ - e (suc m) R.- R.1r ∎ - ℤ↪R (Int.quot m (suc n) i) = e-tr m (suc n) i + ℤ↪R (pos x) = e x + ℤ↪R (negsuc x) = R.- (e (suc x)) open is-ring-hom - ℤ↪R-diff : ∀ m n → ℤ↪R (diff m n) ≡ e m R.- e n - ℤ↪R-diff m zero = R.intror R.inv-unit - ℤ↪R-diff m (suc n) = refl + z-natdiff : ∀ x y → ℤ↪R (x ℕ- y) ≡ e x R.- e y + z-natdiff x zero = R.intror R.inv-unit + z-natdiff zero (suc y) = R.introl refl + z-natdiff (suc x) (suc y) = z-natdiff x y ∙ e-tr x y + + z-add : ∀ x y → ℤ↪R (x +ℤ y) ≡ ℤ↪R x R.+ ℤ↪R y + z-add (pos x) (pos y) = e-add x y + z-add (pos x) (negsuc y) = z-natdiff x (suc y) + z-add (negsuc x) (pos y) = z-natdiff y (suc x) ∙ R.+-commutes + z-add (negsuc x) (negsuc y) = + R.- (e 1 R.+ e (suc x Nat.+ y)) ≡⟨ ap R.-_ (ap₂ R._+_ refl (e-add (suc x) y) ∙ R.extendl R.+-commutes) ⟩ + R.- (e (suc x) R.+ (e 1 R.+ e y)) ≡⟨ R.a.inv-comm ⟩ + (R.- (e 1 R.+ e y)) R.+ (R.- e (suc x)) ≡⟨ R.+-commutes ⟩ + (R.- e (suc x)) R.+ (R.- (e 1 R.+ e y)) ≡⟨ ap₂ R._+_ refl (ap R.-_ (sym (e-add 1 y))) ⟩ + (R.- e (suc x)) R.+ (R.- e (1 Nat.+ y)) ∎ + + z-mul : ∀ x y → ℤ↪R (x *ℤ y) ≡ ℤ↪R x R.* ℤ↪R y + z-mul (pos x) (pos y) = + ℤ↪R (assign pos (x Nat.* y)) ≡⟨ ap ℤ↪R (assign-pos (x Nat.* y)) ⟩ + e (x Nat.* y) ≡⟨ e-mul x y ⟩ + (e x R.* e y) ∎ + z-mul (posz) (negsuc y) = sym R.*-zerol + z-mul (possuc x) (negsuc y) = + R.- e (suc x Nat.* suc y) ≡⟨ ap R.-_ (e-mul (suc x) (suc y)) ⟩ + R.- (e (suc x) R.* e (suc y)) ≡˘⟨ R.*-negater ⟩ + e (suc x) R.* (R.- e (suc y)) ∎ + z-mul (negsuc x) (posz) = + ℤ↪R (assign neg (x Nat.* 0)) ≡⟨ ap ℤ↪R (ap (assign neg) (Nat.*-zeror x)) ⟩ + ℤ↪R 0 ≡⟨ sym R.*-zeror ⟩ + ℤ↪R (negsuc x) R.* R.0r ∎ + z-mul (negsuc x) (possuc y) = + R.- e (suc x Nat.* suc y) ≡⟨ ap R.-_ (e-mul (suc x) (suc y)) ⟩ + R.- (e (suc x) R.* e (suc y)) ≡⟨ sym R.*-negatel ⟩ + (R.- e (suc x)) R.* e (suc y) ∎ + z-mul (negsuc x) (negsuc y) = + e (suc x Nat.* suc y) ≡⟨ e-mul (suc x) (suc y) ⟩ + e (suc x) R.* e (suc y) ≡˘⟨ R.inv-inv ⟩ + R.- (R.- (e (suc x) R.* e (suc y))) ≡˘⟨ ap R.-_ R.*-negater ⟩ + R.- (e (suc x) R.* ℤ↪R (negsuc y)) ≡˘⟨ R.*-negatel ⟩ + ℤ↪R (negsuc x) R.* ℤ↪R (negsuc y) ∎ z→r : Rings.Hom Liftℤ R z→r .hom (lift x) = ℤ↪R x -``` - The last thing we must show is that this is the _unique_ ring homomorphism from the integers to $R$. This, again, is slightly @@ -232,44 +217,14 @@ and that last expression is pretty exactly what our canonical map evaluates to on $n$. So we're done! ```agda - lemma : ∀ (f : Rings.Hom Liftℤ R) i → z→r # lift i ≡ f # lift i - lemma f = - Int-elim-prop (λ _ → hlevel 1) λ a b → sym $ - f # lift (diff a b) ≡⟨ ap (f #_) (ap lift (p a b)) ⟩ - f # lift (diff a 0 +ℤ diff 0 b) ≡⟨ f .preserves .pres-+ (lift (diff a 0)) (lift (diff 0 b)) ⟩ - f # lift (diff a 0) R.+ f # lift (diff 0 b) ≡⟨ ap₂ R._+_ (q a) (is-group-hom.pres-inv gh {x = lift (diff b 0)} ∙ ap R.-_ (q b)) ⟩ - (e a) R.+ (R.- e b) ≡˘⟨ ℤ↪R-diff a b ⟩ - z→r # lift (diff a b) ∎ - where - p : ∀ a b → diff a b ≡ diff a 0 +ℤ diff 0 b - p a b = ap (λ e → diff e b) (sym (Nat.+-zeror a)) - - gh : is-group-hom - (Ring-on.additive-group (Liftℤ .snd) .snd) - (Ring-on.additive-group (R .snd) .snd) - _ - gh = record { pres-⋆ = f .preserves .pres-+ } - - q : ∀ a → f # lift (diff a 0) ≡ e a - q zero = is-group-hom.pres-id gh - q (suc n) = - f # lift (diff (suc n) 0) ≡⟨ f .preserves .pres-+ (lift (diff 1 0)) (lift (diff n 0)) ⟩ - f # lift 1 R.+ f # lift (diff n 0) ≡⟨ ap₂ R._+_ (f .preserves .pres-id) (q n) ⟩ - R.1r R.+ (e n) ≡˘⟨ e-suc n ⟩ - e (suc n) ∎ -``` - -## Abelian groups as Z-modules + module _ (f : Rings.Hom Liftℤ R) where + private module f = is-ring-hom (f .preserves) -A fun consequence of $\ZZ$ being the initial ring is that every -[[abelian group]] admits a unique $\ZZ$-module structure. This is, if -you ask me, rather amazing! The correspondence is as follows: Because of -the delooping-endomorphism ring adjunction, we have a correspondence -between "$R$-module structures on G" and "ring homomorphisms $R \to -\rm{Endo}(G)$" --- and since the latter is contractible, so is the -former! + f-pos : ∀ x → e x ≡ f .hom (lift (pos x)) + f-pos zero = sym f.pres-0 + f-pos (suc x) = e-suc x ∙ sym (f.pres-+ (lift 1) (lift (pos x)) ∙ ap₂ R._+_ f.pres-id (sym (f-pos x))) -```agda -ℤ-module-unique : ∀ (G : Abelian-group ℓ) → is-contr (Ring-action Liftℤ (G .snd)) -ℤ-module-unique G = Equiv→is-hlevel 0 (Action≃Hom Liftℤ G) (Int-is-initial _) + lemma : (i : Int) → z→r # lift i ≡ f # lift i + lemma (pos x) = f-pos x + lemma (negsuc x) = sym (f.pres-neg {lift (possuc x)} ∙ ap R.-_ (sym (f-pos (suc x)))) ``` diff --git a/src/Algebra/Ring/Commutative.lagda.md b/src/Algebra/Ring/Commutative.lagda.md index e9efe38d8..033866bae 100644 --- a/src/Algebra/Ring/Commutative.lagda.md +++ b/src/Algebra/Ring/Commutative.lagda.md @@ -2,8 +2,15 @@ ```agda open import 1Lab.Function.Embedding -open import Algebra.Prelude open import Algebra.Ring + +open import Cat.Displayed.Univalence.Thin +open import Cat.Prelude hiding (_*_ ; _+_) + +open import Data.Int.Properties +open import Data.Int.Base + +import Algebra.Ring.Reasoning as Kit ``` --> @@ -33,7 +40,7 @@ record CRing-on {ℓ} (R : Type ℓ) : Type ℓ where CRing-structure : ∀ ℓ → Thin-structure ℓ CRing-on CRing-structure ℓ = Full-substructure ℓ CRing-on Ring-on emb (Ring-structure ℓ) where - open CRing-on hiding (_↪_) + open CRing-on emb : ∀ X → CRing-on X ↪ Ring-on X emb X .fst = has-ring-on emb X .snd y (r , p) (s , q) = @@ -56,9 +63,17 @@ CRing : ∀ ℓ → Type (lsuc ℓ) CRing ℓ = CRings ℓ .Precategory.Ob module CRing {ℓ} (R : CRing ℓ) where - open CRing-on (R .snd) public + ring : Ring ℓ + ring .fst = R .fst + ring .snd = record { CRing-on (R .snd) } + + open CRing-on (R .snd) using (*-commutes ; _+_ ; _*_ ; -_ ; _-_ ; 1r ; 0r) public + open Kit ring hiding (_+_ ; _*_ ; -_ ; _-_ ; 1r ; 0r) public is-commutative-ring : ∀ {ℓ} (R : Ring ℓ) → Type _ is-commutative-ring R = ∀ {x y} → x R.* y ≡ y R.* x where module R = Ring-on (R .snd) + +ℤ-comm : CRing lzero +ℤ-comm = record { fst = el! Int ; snd = record { has-ring-on = ℤ .snd ; *-commutes = λ {x y} → *ℤ-commutative x y } } ``` diff --git a/src/Algebra/Ring/Ideal.lagda.md b/src/Algebra/Ring/Ideal.lagda.md index d04510500..d5bc464b9 100644 --- a/src/Algebra/Ring/Ideal.lagda.md +++ b/src/Algebra/Ring/Ideal.lagda.md @@ -4,11 +4,16 @@ open import Algebra.Ring.Module.Action open import Algebra.Group.Subgroup open import Algebra.Ring.Module open import Algebra.Group.Ab -open import Algebra.Prelude open import Algebra.Group open import Algebra.Ring +open import Cat.Displayed.Univalence.Thin +open import Cat.Displayed.Total +open import Cat.Prelude + open import Data.Power + +import Algebra.Ring.Reasoning as Ringr ``` --> @@ -42,7 +47,7 @@ multiplication and addition. ```agda module _ {ℓ} (R : Ring ℓ) where - private module R = Ring-on (R .snd) + private module R = Ringr R record is-ideal (𝔞 : ℙ ⌞ R ⌟) : Type (lsuc ℓ) where no-eta-equality @@ -131,15 +136,15 @@ injective. Note that, since our ideals are all two-sided (for simplicity) but our rings may not be commutative (for expressiveness), if we want the ideal -generated by an element to be two-sided, then our ring must be -commutative. +generated by an element to be two-sided, this element must be *central*: +it must commute with every element of the ring. ```agda principal-ideal - : (∀ x y → x R.* y ≡ y R.* x) - → (a : ⌞ R ⌟) + : (a : ⌞ R ⌟) + → (central : ∀ x → a R.* x ≡ x R.* a) → is-ideal λ b → elΩ (Σ _ λ c → b ≡ c R.* a) - principal-ideal comm a = record + principal-ideal a comm = record { has-rep-subgroup = record { has-unit = pure (_ , sym R.*-zerol) ; has-⋆ = λ x y → do @@ -148,15 +153,14 @@ commutative. pure (xi R.+ yi , ap₂ R._+_ p q ∙ sym R.*-distribr) ; has-inv = λ x → do (xi , p) ← x - pure (R.- xi , ap R.-_ p ∙ R.neg-*-l) + pure (R.- xi , ap R.-_ p ∙ sym R.*-negatel) } ; has-*ₗ = λ x y → do (yi , q) ← y - pure (x R.* yi , ap (x R.*_) q ∙ R.*-associative) + pure (x R.* yi , R.m.pushr q) ; has-*ᵣ = λ x y → do (yi , q) ← y pure ( yi R.* x - , ap (R._* x) q ·· sym R.*-associative ·· ap (yi R.*_) (comm a x) - ∙ R.*-associative) + , ap (R._* x) q ∙ R.m.extendr (comm x)) } ``` diff --git a/src/Algebra/Ring/Localisation.lagda.md b/src/Algebra/Ring/Localisation.lagda.md new file mode 100644 index 000000000..4814a97b4 --- /dev/null +++ b/src/Algebra/Ring/Localisation.lagda.md @@ -0,0 +1,431 @@ + + +```agda +module Algebra.Ring.Localisation where +``` + +# Localisations of commutative rings {defines="localisation-of-a-ring"} + +The **localisation** $R\loc{S}$ of a commutative [[ring]] $R$ at a set +of elements $S$ is the universal solution to forcing the elements of $S$ +to become invertible, analogously to how [[localisations of a +category||localisation]] $\cC$ universally invert a class of maps in +$\cC$. Explicitly, it is a commutative ring $R\loc{S}$, equipped with a +homomorphism $-/1 : R \to R\loc{S}$ which sends elements $s \in S$ to +*invertible* elements $s/1 : R\loc{S}$, and which is initial among +these. + +The elements of $R\loc{S}$ are given by formal *fractions*, which we +will write $r/s$, with $r, s : R$ and $s \in S$. Type-theoretically, +such a fraction is really a *triple*, since we must also consider the +witness $w : s \in S$ that the denominator belongs to $S$; but, since +these proofs do not matter for identity of fractions, we will generally +omit them in the prose. + +```agda +record Fraction {ℓ ℓ'} {R : Type ℓ} (S : R → Type ℓ') : Type (ℓ ⊔ ℓ') where + no-eta-equality ; pattern + constructor _/_[_] + field + num : R + denom : R + has-denom : denom ∈ S +``` + + + + + +To ensure that we have a well-behaved theory of fractions, we will +require that $S$ is a **multiplicatively closed** subset of $R$. In +particular, the presence of $1 \in S$ allows us to form the fractions +$r/1$ for any $r : R$, thus ensuring we have a map $R \to R\loc{S}$. We +also require that, if $s \in S$ and $s' \in S$, then also $ss' \in S$. +This will be important to define both multiplication *and identity* of +fractions. + +```agda + record is-multiplicative {ℓ'} (S : ⌞ R ⌟ → Type ℓ') : Type (ℓ ⊔ ℓ') where + field + has-1 : 1r ∈ S + has-* : ∀ {x y} → x ∈ S → y ∈ S → (x * y) ∈ S +``` + +Under the assumption that $S$ is multiplicatively closed, we could +attempt to mimic the usual operations on integer-valued fractions on our +fractions, for example, defining the sum $x/s + y/t$ to be $(xt+ys)/st$. +This turns out not to work too well: for example, if we then also define +$-x/s = (-x)/s$, we would have $x/s + -x/s$ equal to + +$$ +\frac{xs + (-x)s}{s^2} = \frac{(x - x)s}{s^2} = \frac{0}{s^2} +$$ + +which is not *literally* the zero fraction $0/1$. However, we still have +$0s^2 = 0*1$, so these fractions "represent the same quantity" --- they +both stand for *zero times* the formal inverse of something. We could +then try to identify fractions $x/s$ and $y/t$ whenever $xt = ys$, but, +in a general ring $R$, this relation fails to be transitive, so it can +not be equality in a set. Therefore, we allow an "adjustment" by one of +our formal denominators: the equivalence relation we will impose on +fractions identifies $x/s \approx y/t$ whenever there exists $u \in S$ +with $uxt = uys$. + +```agda + data _≈_ {ℓ'} {S : ⌞ R ⌟ → Type ℓ'} (x y : Fraction S) : Type (ℓ ⊔ ℓ') where + inc : (u : ⌞ R ⌟) (u∈S : u ∈ S) → u * ↑ x * ↓ y ≡ u * ↑ y * ↓ x → x ≈ y + squash : is-prop (x ≈ y) +``` + + + +In the literature, this is more commonly phrased as $u(xt - ys) = 0$, +but the equivalence between that and our definition is routine. + +```agda + _≈'_ : ∀ {ℓ'} {S : ⌞ R ⌟ → Type ℓ'} → Fraction S → Fraction S → Type _ + _≈'_ {S = S} (x / s) (y / t) = ∃[ u ∈ R ] (u ∈ S × u * (x * t - y * s) ≡ 0r) +``` + +
+The calculation can be found in this `
`{.html} block.
+ +```agda + ≈→≈' : ∀ {ℓ'} {S : ⌞ R ⌟ → Type ℓ'} {x y : Fraction S} → x ≈ y → x ≈' y + ≈→≈' {x = x / s} {y = y / t} = elim! λ u u∈S p → + let + prf = + u * (x * t - y * s) ≡⟨ solve 5 (λ u x t y s → u :* (x :* t :- y :* s) ≔ u :* x :* t :- u :* y :* s) u x t y s refl ⟩ + u * x * t - u * y * s ≡⟨ ap₂ _-_ refl (sym p) ⟩ + u * x * t - u * x * t ≡⟨ solve 1 (λ x → x :- x ≔ 0) (u * x * t) refl ⟩ + 0r ∎ + in inc (u , u∈S , prf) + + ≈'→≈ : ∀ {ℓ'} {S : ⌞ R ⌟ → Type ℓ'} {x y : Fraction S} → x ≈' y → x ≈ y + ≈'→≈ {x = x / s} {y = y / t} = elim! λ u u∈S p → + let + prf = + u * x * t - u * y * s ≡⟨ solve 5 (λ u x t y s → u :* x :* t :- u :* y :* s ≔ u :* (x :* t :- y :* s)) u x t y s refl ⟩ + u * (x * t - y * s) ≡⟨ p ⟩ + 0r ∎ + in inc u u∈S (zero-diff prf) +``` + +
+ + + +Our next step is showing that this relation is actually an equivalence +relation. The proof of transitivity is the most interesting step: we +assume that $vxt = vys$ and $wyu = wzt$, with "adjustments" by $v, w \in +S$ respectively, but we produce the equation $(wvt)xu = (wvt)zs$ --- we +must consider the denominator of the middle fraction $y/t$ to relate the +two endpoints $x/s$ and $z/u$. + +```agda + Fraction-congruence : Congruence (Fraction (_∈ S)) _ + Fraction-congruence ._∼_ = _≈_ + Fraction-congruence .has-is-prop (_ / _) (_ / _) = hlevel 1 + Fraction-congruence .reflᶜ {x / s} = inc 1r has-1 refl + Fraction-congruence .symᶜ {x / s} {y / t} = rec! λ u u∈S p → inc u u∈S (sym p) + Fraction-congruence ._∙ᶜ_ {x / s} {y / t [ r ]} {z / u} = elim! + λ v v∈S p w w∈S q → + let + prf = + w * v * t * x * u ≡⟨ cring! R ⟩ + w * u * (v * x * t) ≡⟨ ap (w * u *_) p ⟩ + w * u * (v * y * s) ≡⟨ cring! R ⟩ + (w * y * u) * (v * s) ≡⟨ ap₂ _*_ q refl ⟩ + (w * z * t) * (v * s) ≡⟨ cring! R ⟩ + w * v * t * z * s ∎ + in + inc (w * v * t) (has-* (has-* w∈S v∈S) r) prf +``` + + + +We then define $R\loc{S}$ to be the set of fractions $r/s$, identified +according to this relation. Since $1 \in S$, we can immediately cough up +the function $x \mapsto x/1$, mapping $R \to R\loc{S}$. + +```agda + _/1 : ⌞ R ⌟ → Fr.quotient + x /1 = inc (x / 1r [ has-1 ]) +``` + +To define the operations of a ring on $R\loc{S}$, we first define them +at the level of fractions: + +::: mathpar + +$$ +\frac{x}{s} + \frac{y}{t} = \frac{xt + ys}{st} +$$ + +$$ +-\frac{x}{s} = \frac{-x}{s} +$$ + +$$ +\frac{x}{s}\frac{y}{t} = \frac{xy}{st} +$$ + +::: + +As mentioned above, these operations do not satisfy the laws of a ring +*on the set of fractions*. We must therefore prove that they respect the +quotient we've taken, and then prove that they satisfy the ring laws *on +the quotient*. + +```agda + +f : Fraction (_∈ S) → Fraction (_∈ S) → Fraction (_∈ S) + +f (x / s [ p ]) (y / t [ q ]) = (x * t + y * s) / (s * t) [ has-* p q ] + + -f : Fraction (_∈ S) → Fraction (_∈ S) + -f (x / s [ p ]) = (- x) / s [ p ] + + *f : Fraction (_∈ S) → Fraction (_∈ S) → Fraction (_∈ S) + *f (x / s [ p ]) (y / t [ q ]) = (x * y) / (s * t) [ has-* p q ] +``` + +
+Showing that these operations descend to the quotient is +entirely routine algebra; However, the calculations *do* get pretty +long, so we'll leave them in this `
`{.html} block.
+ +```agda + -ₗ_ : Fr.quotient → Fr.quotient + -ₗ_ = Quot-elim (λ _ → hlevel 2) (λ x → inc (-f x)) -f-resp where abstract + -f-resp : ∀ x y → x ≈ y → Path Fr.quotient (inc (-f x)) (inc (-f y)) + -f-resp (x / s) (y / t) = elim! λ u u∈S p → + let + prf = + u * (- x) * t ≡⟨ ap (_* t) *-negater ∙ *-negatel ⟩ + - (u * x * t) ≡⟨ ap -_ p ⟩ + - (u * y * s) ≡⟨ sym *-negatel ∙ ap (_* s) (sym *-negater) ⟩ + u * (- y) * s ∎ + in quot (inc u u∈S prf) + + _+ₗ_ : Fr.quotient → Fr.quotient → Fr.quotient + _+ₗ_ = Fr.op₂-comm +f (λ a b → Fr.reflᶜ' (+f-comm a b)) +f-resp where abstract + +f-comm : ∀ u v → +f u v ≡ +f v u + +f-comm (x / s) (y / t) = Fraction-path +-commutes *-commutes + + +f-resp : ∀ x u v → u ≈ v → +f x u ≈ +f x v + +f-resp (x / s) (u / y) (v / z) = rec! λ w w∈S p → + let + prf = + w * (x * y + u * s) * (s * z) ≡⟨ cring! R ⟩ + w * x * y * s * z + s * s * ⌜ w * u * z ⌝ ≡⟨ ap! p ⟩ + w * x * y * s * z + s * s * ⌜ w * v * y ⌝ ≡⟨ cring! R ⟩ + w * (x * z + v * s) * (s * y) ∎ + in inc w w∈S prf + + _*ₗ_ : Fr.quotient → Fr.quotient → Fr.quotient + _*ₗ_ = Fr.op₂-comm *f *f-comm *f-resp where abstract + *f-comm : ∀ u v → *f u v ≈ *f v u + *f-comm (x / s) (y / t) = inc 1r has-1 (solve 4 (λ x y t s → 1 :* (x :* y) :* (t :* s) ≔ 1 :* (y :* x) :* (s :* t)) x y t s refl) + + *f-resp : ∀ x u v → u ≈ v → *f x u ≈ *f x v + *f-resp (x / s) (u / y) (v / z) = rec! λ w w∈S p → + let + prf = + w * (x * u) * (s * z) ≡⟨ cring! R ⟩ + s * x * (w * u * z) ≡⟨ ap (s * x *_) p ⟩ + s * x * (w * v * y) ≡⟨ cring! R ⟩ + w * (x * v) * (s * y) ∎ + in inc w w∈S prf +``` +
+ +```agda + infixl 8 _+ₗ_ + infixl 9 _*ₗ_ + infix 10 -ₗ_ _/1 + + 0ₗ 1ₗ : Fr.quotient + 0ₗ = 0r /1 + 1ₗ = 1r /1 +``` + +We choose the zero and identity elements for $R\loc{S}$ so that the +localising map $R \to R\loc{S}$ preserves them definitionally. We're +almost done with the construction; while there's quite a bit of algebra +left to do to show that we have a ring, this is almost entirely +automatic. As a warm-up, we can prove that $x/1$ is inverted by $1/x$ +whenever $x \in S$. + +```agda + /1-inverts : ∀ x (p : x ∈ S) → inc (1r / x [ p ]) *ₗ (x /1) ≡ 1ₗ + /1-inverts x p = quot (inc 1r has-1 + (solve 1 (λ x → 1 :* (1 :* x) :* 1 ≔ 1 :* 1 :* (x :* 1)) x refl)) +``` + +The actual proof obligation is shown above: we have to establish that $1 +* (1 * x) * 1 = 1 * 1 * (x * 1)$, which can be shown by a naïve solver. +The equations for each of the ring laws are similarly boring: + +```agda + abstract + +ₗ-idl : ∀ x → 0ₗ +ₗ x ≡ x + +ₗ-idl = elim! λ x s _ → /-ap + (solve 2 (λ s x → 0 :* s :+ x :* 1 ≔ x) s x refl) + *-idl + + +ₗ-invr : ∀ x → x +ₗ (-ₗ x) ≡ 0ₗ + +ₗ-invr = elim! λ x s _ → quot (inc 1r has-1 (solve 2 + (λ x s → 1 :* (x :* s :+ (:- x) :* s) :* 1 ≔ 1 :* 0 :* (s :* s)) x s refl)) + + +ₗ-assoc : ∀ x y z → x +ₗ (y +ₗ z) ≡ (x +ₗ y) +ₗ z + +ₗ-assoc = elim! λ x s _ y t _ z u _ → /-ap + (solve 6 (λ x y z s t u → + x :* (t :* u) :+ (y :* u :+ z :* t) :* s + ≔ (x :* t :+ y :* s) :* u :+ z :* (s :* t)) x y z s t u refl) + *-associative + + +ₗ-comm : ∀ x y → x +ₗ y ≡ y +ₗ x + +ₗ-comm = elim! λ x s _ y t _ → /-ap +-commutes *-commutes + + *ₗ-idl : ∀ x → 1ₗ *ₗ x ≡ x + *ₗ-idl = elim! λ x s s∈S → quot (inc s s∈S (solve 2 + (λ x s → s :* (1 :* x) :* s ≔ s :* x :* (1 :* s)) x s refl)) + + *ₗ-comm : ∀ x y → x *ₗ y ≡ y *ₗ x + *ₗ-comm = elim! λ x s _ y t _ → /-ap *-commutes *-commutes + + *ₗ-assoc : ∀ x y z → x *ₗ (y *ₗ z) ≡ (x *ₗ y) *ₗ z + *ₗ-assoc = elim! λ x s _ y t _ z u _ → /-ap *-associative *-associative + + *ₗ-distribl : ∀ x y z → x *ₗ (y +ₗ z) ≡ (x *ₗ y) +ₗ (x *ₗ z) + *ₗ-distribl = elim! λ x s _ y t _ z u _ → + let + prf : 1r * (x * (y * u + z * t)) * (s * t * (s * u)) + ≡ 1r * (x * y * (s * u) + x * z * (s * t)) * (s * (t * u)) + prf = cring! R + in quot (inc 1r has-1 prf) +``` + +Finally, these fit together to make a commutative ring. + +```agda + private + module mr = make-ring + open make-ring + + mk-loc : make-ring Fr.quotient + mk-loc .ring-is-set = hlevel 2 + mk-loc .0R = 0ₗ + mk-loc ._+_ = _+ₗ_ + mk-loc .-_ = -ₗ_ + mk-loc .1R = 1ₗ + mk-loc ._*_ = _*ₗ_ + mk-loc .+-idl = +ₗ-idl + mk-loc .+-invr = +ₗ-invr + mk-loc .+-assoc = +ₗ-assoc + mk-loc .+-comm = +ₗ-comm + mk-loc .*-idl = *ₗ-idl + mk-loc .*-idr x = *ₗ-comm x 1ₗ ∙ *ₗ-idl x + mk-loc .*-assoc = *ₗ-assoc + mk-loc .*-distribl = *ₗ-distribl + mk-loc .*-distribr x y z = + *ₗ-comm (y +ₗ z) x ∙ *ₗ-distribl x y z ∙ ap₂ _+ₗ_ (*ₗ-comm x y) (*ₗ-comm x z) + + S⁻¹R : CRing ℓ + S⁻¹R .fst = el! Fr.quotient + S⁻¹R .snd .CRing-on.has-ring-on = to-ring-on mk-loc + S⁻¹R .snd .CRing-on.*-commutes {x} {y} = *ₗ-comm x y +``` diff --git a/src/Algebra/Ring/Module/Action.lagda.md b/src/Algebra/Ring/Module/Action.lagda.md index 952e7f871..e91cb6e04 100644 --- a/src/Algebra/Ring/Module/Action.lagda.md +++ b/src/Algebra/Ring/Module/Action.lagda.md @@ -1,5 +1,6 @@ @@ -19,7 +22,7 @@ module Algebra.Ring.Quotient {ℓ} (R : Ring ℓ) where @@ -54,7 +57,7 @@ comprehensibility's sake). private quot-grp : Group _ quot-grp = R.additive-group /ᴳ I.ideal→normal - module R/I = Group-on (quot-grp .snd) + module R/I = Group-on (quot-grp .snd) hiding (magma-hlevel) ``` --> @@ -75,12 +78,12 @@ thing: Since $(x - y) \in I$, also $(xa - ya) \in I$, so $[xa] = [ya]$. ```agda p1 : ∀ a {x y} (r : (x R.- y) ∈ I) → inc (x R.* a) ≡ inc (y R.* a) p1 a {x} {y} x-y∈I = quot $ subst (_∈ I) - (R.*-distribr ∙ ap (x R.* a R.+_) (sym R.neg-*-l)) + (R.*-distribr ∙ ap (x R.* a R.+_) R.*-negatel) (I.has-*ᵣ a x-y∈I) p2 : ∀ a {x y} (r : (x R.- y) ∈ I) → inc (a R.* x) ≡ inc (a R.* y) p2 a {x} {y} x-y∈I = quot $ subst (_∈ I) - (R.*-distribl ∙ ap (a R.* x R.+_) (sym R.neg-*-r)) + (R.*-distribl ∙ ap (a R.* x R.+_) R.*-negater) (I.has-*ₗ a x-y∈I) ``` diff --git a/src/Algebra/Ring/Reasoning.lagda.md b/src/Algebra/Ring/Reasoning.lagda.md new file mode 100644 index 000000000..34c90751c --- /dev/null +++ b/src/Algebra/Ring/Reasoning.lagda.md @@ -0,0 +1,51 @@ + + +```agda +module Algebra.Ring.Reasoning {ℓ} (R : Ring ℓ) where +``` + +# Reasoning combinators for rings + + + +```agda +*-zerol : 0r * x ≡ 0r +*-zerol {x} = + 0r * x ≡⟨ a.introl a.inversel ⟩ + (- 0r * x) + 0r * x + 0r * x ≡⟨ a.pullr (sym *-distribr) ⟩ + (- 0r * x) + (0r + 0r) * x ≡⟨ ap₂ _+_ refl (ap (_* x) a.idl) ⟩ + (- 0r * x) + 0r * x ≡⟨ a.inversel ⟩ + 0r ∎ + +*-zeror : x * 0r ≡ 0r +*-zeror {x} = + x * 0r ≡⟨ a.intror a.inverser ⟩ + x * 0r + (x * 0r - x * 0r) ≡⟨ a.pulll (sym *-distribl) ⟩ + x * (0r + 0r) - x * 0r ≡⟨ ap₂ _-_ (ap (x *_) a.idl) refl ⟩ + x * 0r - x * 0r ≡⟨ a.inverser ⟩ + 0r ∎ + +*-negatel : (- x) * y ≡ - (x * y) +*-negatel {x} {y} = monoid-inverse-unique a.has-is-monoid (x * y) ((- x) * y) (- (x * y)) + (sym *-distribr ·· ap (_* y) a.inversel ·· *-zerol) + a.inverser + +*-negater : x * (- y) ≡ - (x * y) +*-negater {x} {y} = monoid-inverse-unique a.has-is-monoid (x * y) (x * (- y)) (- (x * y)) + (sym *-distribl ·· ap (x *_) a.inversel ·· *-zeror) + a.inverser +``` diff --git a/src/Algebra/Ring/Solver.agda b/src/Algebra/Ring/Solver.agda index bf1107915..8999382b8 100644 --- a/src/Algebra/Ring/Solver.agda +++ b/src/Algebra/Ring/Solver.agda @@ -16,25 +16,40 @@ open import 1Lab.Reflection open import Algebra.Ring.Cat.Initial open import Algebra.Ring.Commutative open import Algebra.Group.Ab -open import Algebra.Prelude open import Algebra.Group open import Algebra.Ring +open import Cat.Displayed.Total +open import Cat.Prelude hiding (_+_ ; _*_ ; _-_) + +open import Data.Fin.Product open import Data.Fin.Base -open import Data.Int.HIT -open import Data.List hiding (lookup) +open import Data.Int.Base +open import Data.List hiding (lookup ; tabulate) open import Data.Dec open import Data.Nat +import Algebra.Ring.Reasoning as Kit + +import Data.Int.Base as B + +open Total-hom + module Algebra.Ring.Solver where module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where private - module R = CRing-on cring - ℤ↪R-rh = Int-is-initial (el _ R.has-is-set , R.has-ring-on) .centre + R' : Ring _ + R' = record { fst = el _ (CRing-on.has-is-set cring) ; snd = CRing-on.has-ring-on cring } + + module R = Kit R' + + ℤ↪R-rh = Int-is-initial R' .centre module ℤ↪R = is-ring-hom (ℤ↪R-rh .preserves) embed-coe = ℤ↪R-rh .hom + open CRing-on cring using (*-commutes) + data Poly : Nat → Type ℓ data Normal : Nat → Type ℓ @@ -137,7 +152,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where -ₙ_ : ∀ {n} → Normal n → Normal n -ₚ_ : ∀ {n} → Poly (suc n) → Poly (suc n) - -ₙ con x = con (negate x) + -ₙ con x = con (negℤ x) -ₙ poly x = poly (-ₚ x) -ₚ x = (-ₙ 1n) *ₙₚ x @@ -154,8 +169,8 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where -- Short-hand notation. - infix 40 _:-_ - infix 30 _:*_ + infixl 30 _:-_ _:+_ + infixl 40 _:*_ _:+_ : ∀ {n} → Polynomial n → Polynomial n → Polynomial n _:+_ = op [+] @@ -180,6 +195,10 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where ⟦⟧-Polynomial : ∀ {n} → ⟦⟧-notation (Polynomial n) ⟦⟧-Polynomial = brackets _ eval + Number-Polynomial : ∀ {n} → Number (Polynomial n) + Number-Polynomial .Number.Constraint x = Lift _ ⊤ + Number-Polynomial .Number.fromNat n = con (pos n) + eval (op o p₁ p₂) ρ = sem o (⟦ p₁ ⟧ ρ) (⟦ p₂ ⟧ ρ) eval (con c) ρ = embed-coe (lift c) eval (var x) ρ = lookup ρ x @@ -307,7 +326,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where where lem₁' = a R.* c R.* x ≡˘⟨ R.*-associative ⟩ - a R.* ⌜ c R.* x ⌝ ≡⟨ ap! R.*-commutes ⟩ + a R.* ⌜ c R.* x ⌝ ≡⟨ ap! *-commutes ⟩ a R.* (x R.* c) ≡⟨ R.*-associative ⟩ a R.* x R.* c ∎ @@ -319,7 +338,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where lem₂ = (a R.* d R.+ b R.* c) R.* x ≡⟨ R.*-distribr ⟩ a R.* d R.* x R.+ b R.* c R.* x ≡˘⟨ ap₂ R._+_ R.*-associative R.*-associative ⟩ - a R.* ⌜ d R.* x ⌝ R.+ b R.* (c R.* x) ≡⟨ ap! R.*-commutes ⟩ + a R.* ⌜ d R.* x ⌝ R.+ b R.* (c R.* x) ≡⟨ ap! *-commutes ⟩ a R.* (x R.* d) R.+ b R.* (c R.* x) ≡⟨ ap₂ R._+_ R.*-associative refl ⟩ a R.* x R.* d R.+ b R.* (c R.* x) ∎ @@ -337,7 +356,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where *ₚₙ-hom c (p *x+ d) x ρ with c ==ₙ 0n ... | just c=0 = sym (ap₂ R._*_ refl (ap (λ e → En e ρ) c=0 ∙ 0n-hom ρ) ∙ R.*-zeror) ... | nothing = - ap₂ R._+_ (ap (R._* x) (*ₚₙ-hom c p x ρ) ·· sym R.*-associative ·· ap₂ R._*_ refl R.*-commutes ∙ R.*-associative) + ap₂ R._+_ (ap (R._* x) (*ₚₙ-hom c p x ρ) ·· sym R.*-associative ·· ap₂ R._*_ refl *-commutes ∙ R.*-associative) (*ₙ-hom d c ρ) ∙ sym R.*-distribr @@ -347,7 +366,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where -ₚ-hom p (x ∷ ρ) = *ₙₚ-hom (-ₙ 1n) p x ρ ∙ ap₂ R._*_ (-ₙ-hom 1n ρ ∙ ap R.-_ (1n-hom ρ)) refl - ∙ sym R.neg-*-l ∙ ap R.-_ R.*-idl + ∙ R.*-negatel ∙ ap R.-_ R.*-idl -ₙ-hom (con x) ρ = ℤ↪R.pres-neg {x = lift x} -ₙ-hom (poly x) ρ = -ₚ-hom x ρ @@ -388,6 +407,30 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where test-identities x = solve (var 0 :+ (con 0 :* con 1)) ((con 1 :+ con 0) :* var 0) (x ∷ []) refl +module Explicit {ℓ} (R : CRing ℓ) where + private module I = Impl (R .snd) + + open I renaming (solve to solve-impl) + open I public using (Polynomial ; _:+_ ; _:-_ ; :-_ ; _:*_ ; con ; Number-Polynomial) + + _≔_ : ∀ {n} → Polynomial n → Polynomial n → Polynomial n × Polynomial n + x ≔ y = x , y + + private + variables : ∀ {n} → Πᶠ {n = n} λ i → Polynomial n + variables = tabulateₚ var + + abstract + solve + : (n : Nat) (f : Arrᶠ {n = n} (λ i → Polynomial n) (Polynomial n × Polynomial n)) + → (let (lhs , rhs) = applyᶠ {n = n} f variables) + → ∀ᶠ n (λ i → ⌞ R ⌟) λ vs + → (let rs = tabulate (indexₚ vs)) + → En (normal lhs) rs ≡ En (normal rhs) rs + → ⟦ lhs ⟧ rs ≡ ⟦ rhs ⟧ rs + solve n f = curry-∀ᶠ {n = n} (λ a → solve-impl lhs rhs (tabulate (indexₚ a))) + where open Σ (applyᶠ {n = n} f variables) renaming (fst to lhs ; snd to rhs) + module Reflection where private pattern ring-args cring args = (_ hm∷ _ hm∷ cring v∷ args) @@ -421,11 +464,11 @@ module Reflection where build-expr : ∀ {ℓ} {A : Type ℓ} → Term → Variables A → Term → TC (Term × Variables A) build-expr cring vs (“0” cring') = do unify cring cring' - z ← quoteTC (diff 0 0) + z ← quoteTC (Int.pos 0) pure $ con (quote Impl.Polynomial.con) (z v∷ []) , vs build-expr cring vs (“1” cring') = do unify cring cring' - o ← quoteTC (diff 1 0) + o ← quoteTC (Int.pos 1) pure $ con (quote Impl.Polynomial.con) (o v∷ []) , vs build-expr cring vs (“*” cring' t1 t2) = do unify cring cring' @@ -496,3 +539,14 @@ private module TestCRing {ℓ} (R : CRing ℓ) where test-identities : ∀ x → x R.+ (R.0r R.* R.1r) ≡ (R.1r R.+ R.0r) R.* x test-identities x = cring! R + + test-negation : ∀ x y → x R.* (R.- y) ≡ R.- (x R.* y) + test-negation x y = cring! R + +private module TestExplicit where + open Explicit ℤ-comm + + _ : ∀ x y u v → (x B.*ℤ y) B.*ℤ (u B.*ℤ v) ≡ (x B.*ℤ u) B.*ℤ (y B.*ℤ v) + _ = λ x y u v → solve 4 + (λ x y u v → (x :* y) :* (u :* v) ≔ (x :* u) :* (y :* v)) + x y u v refl diff --git a/src/Cat/Abelian/Base.lagda.md b/src/Cat/Abelian/Base.lagda.md index 8ba3339bf..a6f364d38 100644 --- a/src/Cat/Abelian/Base.lagda.md +++ b/src/Cat/Abelian/Base.lagda.md @@ -2,8 +2,7 @@ ```agda open import Algebra.Group.Ab.Tensor open import Algebra.Group.Ab -open import Algebra.Prelude -open import Algebra.Monoid hiding (idl; idr) +open import Algebra.Monoid open import Algebra.Group open import Cat.Diagram.Equaliser.Kernel @@ -12,12 +11,17 @@ open import Cat.Diagram.Biproduct open import Cat.Diagram.Coproduct open import Cat.Diagram.Terminal open import Cat.Diagram.Product +open import Cat.Displayed.Total +open import Cat.Instances.Slice open import Cat.Diagram.Zero +open import Cat.Prelude hiding (_+_ ; _*_ ; _-_) import Algebra.Group.Cat.Base as Grp import Algebra.Group.Ab.Hom as Ab -import Cat.Reasoning +import Cat.Reasoning as Cat + +open Total-hom ``` --> @@ -114,17 +118,17 @@ $-ab = (-a)b = a(-b)$, etc. Hom.inverse (0m ∘ f) + (0m ∘ f) ≡⟨ Hom.inversel ⟩ 0m ∎ - neg-∘-l + ∘-negatel : ∀ {A B C} {g : Hom B C} {h : Hom A B} → Hom.inverse (g ∘ h) ≡ Hom.inverse g ∘ h - neg-∘-l {g = g} {h} = monoid-inverse-unique Hom.has-is-monoid (g ∘ h) _ _ + ∘-negatel {g = g} {h} = monoid-inverse-unique Hom.has-is-monoid (g ∘ h) _ _ Hom.inversel (∘-linear-l _ _ _ ∙ ap (_∘ h) Hom.inverser ∙ ∘-zero-l) - neg-∘-r + ∘-negater : ∀ {A B C} {g : Hom B C} {h : Hom A B} → Hom.inverse (g ∘ h) ≡ g ∘ Hom.inverse h - neg-∘-r {g = g} {h} = monoid-inverse-unique Hom.has-is-monoid (g ∘ h) _ _ + ∘-negater {g = g} {h} = monoid-inverse-unique Hom.has-is-monoid (g ∘ h) _ _ Hom.inversel (∘-linear-r _ _ _ ∙ ap (g ∘_) Hom.inverser ∙ ∘-zero-r) @@ -132,7 +136,7 @@ $-ab = (-a)b = a(-b)$, etc. : ∀ {A B C} (f g : Hom B C) (h : Hom A B) → (f ∘ h) - (g ∘ h) ≡ (f - g) ∘ h ∘-minus-l f g h = - f ∘ h - g ∘ h ≡⟨ ap (f ∘ h +_) neg-∘-l ⟩ + f ∘ h - g ∘ h ≡⟨ ap (f ∘ h +_) ∘-negatel ⟩ f ∘ h + (Hom.inverse g ∘ h) ≡⟨ ∘-linear-l _ _ _ ⟩ (f - g) ∘ h ∎ @@ -140,7 +144,7 @@ $-ab = (-a)b = a(-b)$, etc. : ∀ {A B C} (f : Hom B C) (g h : Hom A B) → (f ∘ g) - (f ∘ h) ≡ f ∘ (g - h) ∘-minus-r f g h = - f ∘ g - f ∘ h ≡⟨ ap (f ∘ g +_) neg-∘-r ⟩ + f ∘ g - f ∘ h ≡⟨ ap (f ∘ g +_) ∘-negater ⟩ f ∘ g + (f ∘ Hom.inverse h) ≡⟨ ∘-linear-r _ _ _ ⟩ f ∘ (g - h) ∎ ``` @@ -333,7 +337,7 @@ each $\hom$-monoid becomes a $\hom$-*group*. diff --git a/src/Cat/Abelian/Instances/Functor.lagda.md b/src/Cat/Abelian/Instances/Functor.lagda.md index 30519aa53..9a3524738 100644 --- a/src/Cat/Abelian/Instances/Functor.lagda.md +++ b/src/Cat/Abelian/Instances/Functor.lagda.md @@ -97,9 +97,9 @@ natural. grp .inv f .η x = B.Hom.inverse (f .η x) grp .inv f .is-natural x y g = - B.Hom.inverse (f .η y) B.∘ F.₁ g ≡˘⟨ B.neg-∘-l ⟩ + B.Hom.inverse (f .η y) B.∘ F.₁ g ≡˘⟨ B.∘-negatel ⟩ B.Hom.inverse ⌜ f .η y B.∘ F.₁ g ⌝ ≡⟨ ap! (f .is-natural x y g) ⟩ - B.Hom.inverse (G.₁ g B.∘ f .η x) ≡⟨ B.neg-∘-r ⟩ + B.Hom.inverse (G.₁ g B.∘ f .η x) ≡⟨ B.∘-negater ⟩ G.₁ g B.∘ B.Hom.inverse (f .η x) ∎ grp .assoc _ _ _ = ext λ _ → B.Hom.associative diff --git a/src/Cat/CartesianClosed/Instances/PSh.agda b/src/Cat/CartesianClosed/Instances/PSh.agda index ebf830d35..84fbf92d5 100644 --- a/src/Cat/CartesianClosed/Instances/PSh.agda +++ b/src/Cat/CartesianClosed/Instances/PSh.agda @@ -1,8 +1,12 @@ open import Cat.Instances.Functor.Limits open import Cat.Instances.Sets.Complete +open import Cat.Diagram.Coequaliser open import Cat.Diagram.Exponential -open import Cat.Diagram.Everything +open import Cat.Diagram.Coproduct open import Cat.Instances.Functor +open import Cat.Diagram.Pullback +open import Cat.Diagram.Terminal +open import Cat.Diagram.Product open import Cat.Functor.Adjoint open import Cat.Instances.Sets open import Cat.Functor.Hom diff --git a/src/Cat/Diagram/Coproduct/Indexed.lagda.md b/src/Cat/Diagram/Coproduct/Indexed.lagda.md index 5e0b4cbee..369cd971d 100644 --- a/src/Cat/Diagram/Coproduct/Indexed.lagda.md +++ b/src/Cat/Diagram/Coproduct/Indexed.lagda.md @@ -92,7 +92,7 @@ Indexed-coproduct-≃ e {F} p = λ where Lift-Indexed-coproduct : ∀ {ℓ} ℓ' → {I : Type ℓ} → {F : I → C.Ob} - → Indexed-coproduct {Idx = Lift ℓ' I} (F ⊙ Lift.lower) + → Indexed-coproduct {Idx = Lift ℓ' I} (F ⊙ lower) → Indexed-coproduct F Lift-Indexed-coproduct _ = Indexed-coproduct-≃ (Lift-≃ e⁻¹) diff --git a/src/Cat/Diagram/Everything.agda b/src/Cat/Diagram/Everything.agda deleted file mode 100644 index d0cb929b9..000000000 --- a/src/Cat/Diagram/Everything.agda +++ /dev/null @@ -1,30 +0,0 @@ -module Cat.Diagram.Everything where -open import Cat.Diagram.Coequaliser public -open import Cat.Diagram.Coequaliser.RegularEpi public -open import Cat.Diagram.Colimit.Base public -open import Cat.Diagram.Congruence public -open import Cat.Diagram.Coproduct.Indexed public -open import Cat.Diagram.Coproduct public -open import Cat.Diagram.Duals public -open import Cat.Diagram.Equaliser.Kernel public -open import Cat.Diagram.Equaliser public -open import Cat.Diagram.Equaliser.RegularMono public -open import Cat.Diagram.Idempotent public -open import Cat.Diagram.Image public -open import Cat.Diagram.Initial public -open import Cat.Diagram.Limit.Base public -open import Cat.Diagram.Limit.Equaliser public -open import Cat.Diagram.Limit.Finite public -open import Cat.Diagram.Limit.Product public -open import Cat.Diagram.Limit.Pullback public -open import Cat.Diagram.Monad public -open import Cat.Diagram.Monad.Limits public -open import Cat.Diagram.Product.Indexed public -open import Cat.Diagram.Product public -open import Cat.Diagram.Pullback public -open import Cat.Diagram.Pullback.Properties public -open import Cat.Diagram.Pushout public -open import Cat.Diagram.Pushout.Properties public -open import Cat.Diagram.Sieve public -open import Cat.Diagram.Terminal public -open import Cat.Diagram.Zero public diff --git a/src/Cat/Diagram/Product/Indexed.lagda.md b/src/Cat/Diagram/Product/Indexed.lagda.md index bad06d9ab..acd51be3a 100644 --- a/src/Cat/Diagram/Product/Indexed.lagda.md +++ b/src/Cat/Diagram/Product/Indexed.lagda.md @@ -103,7 +103,7 @@ Indexed-product-≃ e {F} p = λ where Lift-Indexed-product : ∀ {ℓ} ℓ' → {I : Type ℓ} → {F : I → C.Ob} - → Indexed-product {Idx = Lift ℓ' I} (F ⊙ Lift.lower) + → Indexed-product {Idx = Lift ℓ' I} (F ⊙ lower) → Indexed-product F Lift-Indexed-product _ = Indexed-product-≃ (Lift-≃ e⁻¹) diff --git a/src/Cat/Displayed/Cartesian/Indexing.lagda.md b/src/Cat/Displayed/Cartesian/Indexing.lagda.md index a650cc837..d691ab1f7 100644 --- a/src/Cat/Displayed/Cartesian/Indexing.lagda.md +++ b/src/Cat/Displayed/Cartesian/Indexing.lagda.md @@ -1,6 +1,5 @@ @@ -260,12 +264,12 @@ into an identification. F : Displayed-functor Mon[ Setsₓ ] Mon Id F .F₀' o .identity = o .η (lift tt) F .F₀' o ._⋆_ x y = o .μ (x , y) - F .F₀' o .has-is-monoid .Mon.has-is-semigroup = + F .F₀' o .has-is-monoid .has-is-semigroup = record { has-is-magma = record { has-is-set = hlevel 2 } ; associative = o .μ-assoc $ₚ _ } - F .F₀' o .has-is-monoid .Mon.idl = o .μ-unitl $ₚ _ - F .F₀' o .has-is-monoid .Mon.idr = o .μ-unitr $ₚ _ + F .F₀' o .has-is-monoid .idl = o .μ-unitl $ₚ _ + F .F₀' o .has-is-monoid .idr = o .μ-unitr $ₚ _ F .F₁' wit .pres-id = wit .pres-η $ₚ _ F .F₁' wit .pres-⋆ x y = wit .pres-μ $ₚ _ F .F-id' = prop! diff --git a/src/Cat/Monoidal/Diagram/Monoid/Representable.lagda.md b/src/Cat/Monoidal/Diagram/Monoid/Representable.lagda.md index acd55af48..370ca9a92 100644 --- a/src/Cat/Monoidal/Diagram/Monoid/Representable.lagda.md +++ b/src/Cat/Monoidal/Diagram/Monoid/Representable.lagda.md @@ -2,7 +2,7 @@ ```agda open import Algebra.Monoid.Category open import Algebra.Semigroup -open import Algebra.Monoid renaming (idr to mon-idr; idl to mon-idl) +open import Algebra.Monoid open import Algebra.Magma open import Cat.Monoidal.Instances.Cartesian @@ -22,6 +22,8 @@ open import Cat.Functor.Hom open import Cat.Prelude import Cat.Reasoning + +open is-monoid renaming (idl to mon-idl ; idr to mon-idr) ``` --> diff --git a/src/Data/Dec/Base.lagda.md b/src/Data/Dec/Base.lagda.md index 79b0b0b8a..e4bf2476b 100644 --- a/src/Data/Dec/Base.lagda.md +++ b/src/Data/Dec/Base.lagda.md @@ -224,3 +224,11 @@ from-dec-is-equiv = is-iso→is-equiv (iso to-dec p q) where q (yes x) = refl q (no x) = refl ``` + + diff --git a/src/Data/Fin/Closure.lagda.md b/src/Data/Fin/Closure.lagda.md index 126de3e2f..12b460aaf 100644 --- a/src/Data/Fin/Closure.lagda.md +++ b/src/Data/Fin/Closure.lagda.md @@ -4,8 +4,8 @@ open import 1Lab.Prelude open import Data.Set.Coequaliser open import Data.Fin.Properties -open import Data.Nat.Prime open import Data.Fin.Base +open import Data.Nat.Base open import Data.Dec open import Data.Sum diff --git a/src/Data/Fin/Product.lagda.md b/src/Data/Fin/Product.lagda.md index e5cdb90e3..a0d1403c7 100644 --- a/src/Data/Fin/Product.lagda.md +++ b/src/Data/Fin/Product.lagda.md @@ -57,6 +57,17 @@ tabulateₚ {n = zero} f = tt tabulateₚ {n = suc n} f = f fzero , tabulateₚ λ i → f (fsuc i) ``` + + Elements of $\Pi^f$ for sequences with a known length enjoy strong extensionality properties, since they are iterated types with $\eta$-expansion. As an example: @@ -108,6 +119,25 @@ mapₚ {0} f xs = xs mapₚ {suc n} f xs = f fzero (xs .fst) , mapₚ (λ i → f (fsuc i)) (xs .snd) ``` + + More generically, we can characterise the entries of an updated product type. @@ -145,6 +175,27 @@ Arrᶠ {0} P x = x Arrᶠ {suc n} P x = P fzero → Arrᶠ (λ i → P (fsuc i)) x ``` + + In the generic case, a finitary curried function can be eliminated using a finitary dependent product; Moreover, curried functions are "extensional" with respect to this application. diff --git a/src/Data/Int/Base.lagda.md b/src/Data/Int/Base.lagda.md index 157aa2073..771c9af79 100644 --- a/src/Data/Int/Base.lagda.md +++ b/src/Data/Int/Base.lagda.md @@ -2,12 +2,13 @@ ``` open import 1Lab.Path.IdentitySystem open import 1Lab.HLevel.Closure +open import 1Lab.HLevel open import 1Lab.Equiv open import 1Lab.Path open import 1Lab.Type open import Data.Dec.Base -open import Data.Nat.Base +open import Data.Nat.Base hiding (Positive) ``` --> @@ -303,7 +304,8 @@ sign× pos neg = neg sign× neg pos = neg _*ℤ_ : Int → Int → Int -i *ℤ j = assign (sign× (sign i) (sign j)) (abs i * abs j) +i@(pos _) *ℤ j = assign (sign× pos (sign j)) (abs i * abs j) +i@(negsuc _) *ℤ j = assign (sign× neg (sign j)) (abs i * abs j) ``` There are actually alternative definitions of addition and @@ -346,3 +348,38 @@ minℤ (pos _) (negsuc y) = negsuc y minℤ (negsuc x) (pos _) = negsuc x minℤ (negsuc x) (negsuc y) = negsuc (max x y) ``` + + diff --git a/src/Data/Int/DivMod.lagda.md b/src/Data/Int/DivMod.lagda.md index a6fa2f9ff..ab758f8ad 100644 --- a/src/Data/Int/DivMod.lagda.md +++ b/src/Data/Int/DivMod.lagda.md @@ -10,7 +10,7 @@ open import Data.Int.Divisible open import Data.Nat.DivMod open import Data.Dec.Base open import Data.Fin hiding (_<_) -open import Data.Int +open import Data.Int hiding (Positive) open import Data.Nat as Nat ``` --> @@ -85,7 +85,7 @@ $b - (-a)\%b$ as the remainder. lemma = assign neg (q * b + suc r) ≡⟨ ap (assign neg) (+-commutative (q * b) _) ⟩ negsuc (r + q * b) ≡˘⟨ negℤ-+ℤ-negsuc r (q * b) ⟩ - negℤ (pos r) +ℤ negsuc (q * b) ≡⟨ ap (_+ℤ negsuc (q * b)) (ℤ.insertl {h = negℤ (pos b')} (+ℤ-invl (pos b')) {f = negℤ (pos r)}) ⟩ + negℤ (pos r) +ℤ negsuc (q * b) ≡⟨ ap (_+ℤ negsuc (q * b)) (ℤ.insertl {negℤ (pos b')} (+ℤ-invl (pos b')) {negℤ (pos r)}) ⟩ ⌜ negℤ (pos b') ⌝ +ℤ (pos b' -ℤ pos r) +ℤ negsuc (q * b) ≡˘⟨ ap¡ (assign-neg b') ⟩ assign neg b' +ℤ (pos b' -ℤ pos r) +ℤ negsuc (q * b) ≡⟨ ap (_+ℤ negsuc (q * b)) (+ℤ-commutative (assign neg b') (pos b' -ℤ pos r)) ⟩ (pos b' -ℤ pos r) +ℤ assign neg b' +ℤ negsuc (q * b) ≡˘⟨ +ℤ-assoc (pos b' -ℤ pos r) _ _ ⟩ @@ -186,7 +186,7 @@ divides-diff→same-rem n x y n∣x-y p' : y ≡ (q -ℤ k) *ℤ pos n +ℤ pos r p' = y ≡˘⟨ negℤ-negℤ y ⟩ - negℤ (negℤ y) ≡⟨ ℤ.insertr (ℤ.inversel {x}) {f = negℤ (negℤ y)} ⟩ + negℤ (negℤ y) ≡⟨ ℤ.insertr (ℤ.inversel {x}) {negℤ (negℤ y)} ⟩ ⌜ negℤ (negℤ y) +ℤ negℤ x ⌝ +ℤ x ≡⟨ ap! (+ℤ-commutative (negℤ (negℤ y)) _) ⟩ ⌜ negℤ x -ℤ negℤ y ⌝ +ℤ x ≡˘⟨ ap¡ (negℤ-distrib x (negℤ y)) ⟩ negℤ ⌜ x -ℤ y ⌝ +ℤ x ≡˘⟨ ap¡ z ⟩ diff --git a/src/Data/Int/Order.lagda.md b/src/Data/Int/Order.lagda.md index c6cd5b2c9..2a101b4d2 100644 --- a/src/Data/Int/Order.lagda.md +++ b/src/Data/Int/Order.lagda.md @@ -57,6 +57,12 @@ for the ordering on natural numbers. ≤-refl {x = pos x} = pos≤pos Nat.≤-refl ≤-refl {x = negsuc x} = neg≤neg Nat.≤-refl +≤-refl' : ∀ {x y} → x ≡ y → x ≤ y +≤-refl' {pos x} {pos y} p = pos≤pos (Nat.≤-refl' (pos-injective p)) +≤-refl' {pos x} {negsuc y} p = absurd (pos≠negsuc p) +≤-refl' {negsuc x} {pos y} p = absurd (negsuc≠pos p) +≤-refl' {negsuc x} {negsuc y} p = neg≤neg (Nat.≤-refl' (negsuc-injective (sym p))) + ≤-trans : ∀ {x y z} → x ≤ y → y ≤ z → x ≤ z ≤-trans (neg≤neg p) (neg≤neg q) = neg≤neg (Nat.≤-trans q p) ≤-trans (neg≤neg p) neg≤pos = neg≤pos @@ -66,6 +72,18 @@ for the ordering on natural numbers. ≤-antisym : ∀ {x y} → x ≤ y → y ≤ x → x ≡ y ≤-antisym (neg≤neg p) (neg≤neg q) = ap negsuc (Nat.≤-antisym q p) ≤-antisym (pos≤pos p) (pos≤pos q) = ap pos (Nat.≤-antisym p q) + +unpos≤pos : ∀ {x y} → pos x ≤ pos y → x Nat.≤ y +unpos≤pos (pos≤pos p) = p + +unneg≤neg : ∀ {x y} → negsuc x ≤ negsuc y → y Nat.≤ x +unneg≤neg (neg≤neg p) = p + +apos≤apos : ∀ {x y} → x Nat.≤ y → assign pos x ≤ assign pos y +apos≤apos {x} {y} p = ≤-trans (≤-refl' (assign-pos x)) (≤-trans (pos≤pos p) (≤-refl' (sym (assign-pos y)))) + +unapos≤apos : ∀ {x y} → assign pos x ≤ assign pos y → x Nat.≤ y +unapos≤apos {x} {y} p = unpos≤pos (≤-trans (≤-refl' (sym (assign-pos x))) (≤-trans p (≤-refl' (assign-pos y)))) ``` ## Totality @@ -187,15 +205,15 @@ rotℤ≤l (negsuc zero) x y p = pred-≤ _ _ p rotℤ≤l (negsuc (suc k)) x y p = pred-≤ _ _ (rotℤ≤l (negsuc k) x y p) abstract - +ℤ-mono-l : ∀ k x y → x ≤ y → (k +ℤ x) ≤ (k +ℤ y) - +ℤ-mono-l k x y p = transport + +ℤ-preserves-≤l : ∀ k x y → x ≤ y → (k +ℤ x) ≤ (k +ℤ y) + +ℤ-preserves-≤l k x y p = transport (sym (ap₂ _≤_ (rot-is-add k x) (rot-is-add k y))) (rotℤ≤l k x y p) - +ℤ-mono-r : ∀ k x y → x ≤ y → (x +ℤ k) ≤ (y +ℤ k) - +ℤ-mono-r k x y p = transport + +ℤ-preserves-≤r : ∀ k x y → x ≤ y → (x +ℤ k) ≤ (y +ℤ k) + +ℤ-preserves-≤r k x y p = transport (ap₂ _≤_ (+ℤ-commutative k x) (+ℤ-commutative k y)) - (+ℤ-mono-l k x y p) + (+ℤ-preserves-≤l k x y p) negℤ-anti : ∀ x y → x ≤ y → negℤ y ≤ negℤ x negℤ-anti posz posz x≤y = x≤y @@ -211,4 +229,20 @@ abstract negℤ-anti-full (possuc x) (possuc y) (neg≤neg x≤y) = pos≤pos (Nat.s≤s x≤y) negℤ-anti-full (negsuc x) (pos y) _ = neg≤pos negℤ-anti-full (negsuc x) (negsuc y) (pos≤pos (Nat.s≤s y≤x)) = neg≤neg y≤x + + *ℤ-cancel-≤r : ∀ {x y z} ⦃ _ : Positive x ⦄ → (y *ℤ x) ≤ (z *ℤ x) → y ≤ z + *ℤ-cancel-≤r {possuc x} {y = pos y} {pos z} ⦃ pos _ ⦄ p = pos≤pos + (Nat.*-cancel-≤r (suc x) (unapos≤apos p)) + *ℤ-cancel-≤r {possuc x} {y = pos y} {negsuc z} ⦃ pos _ ⦄ p = absurd (¬pos≤neg (≤-trans (≤-refl' (sym (assign-pos (y * suc x)))) p)) + *ℤ-cancel-≤r {possuc x} {y = negsuc y} {pos z} ⦃ pos _ ⦄ p = neg≤pos + *ℤ-cancel-≤r {possuc x} {y = negsuc y} {negsuc z} ⦃ pos _ ⦄ p = neg≤neg + (Nat.*-cancel-≤r (suc x) (Nat.+-reflects-≤l (z * suc x) (y * suc x) x (unneg≤neg p))) + + *ℤ-preserves-≤r : ∀ {x y} z ⦃ _ : Positive z ⦄ → x ≤ y → (x *ℤ z) ≤ (y *ℤ z) + *ℤ-preserves-≤r {pos x} {pos y} (possuc z) ⦃ pos z ⦄ p = apos≤apos {x * suc z} {y * suc z} (Nat.*-preserves-≤r x y (suc z) (unpos≤pos p)) + *ℤ-preserves-≤r {negsuc x} {pos y} (possuc z) ⦃ pos z ⦄ p = ≤-trans neg≤pos (≤-refl' (sym (assign-pos (y * suc z)))) + *ℤ-preserves-≤r {negsuc x} {negsuc y} (possuc z) ⦃ pos z ⦄ p = neg≤neg (Nat.+-preserves-≤l (y * suc z) (x * suc z) z (Nat.*-preserves-≤r y x (suc z) (unneg≤neg p))) + + *ℤ-nonnegative : ∀ {x y} → 0 ≤ x → 0 ≤ y → 0 ≤ (x *ℤ y) + *ℤ-nonnegative {pos x} {pos y} (pos≤pos p) (pos≤pos q) = ≤-trans (pos≤pos Nat.0≤x) (≤-refl' (sym (assign-pos (x * y)))) ``` diff --git a/src/Data/Int/Properties.lagda.md b/src/Data/Int/Properties.lagda.md index ced6cc3c2..08985f55e 100644 --- a/src/Data/Int/Properties.lagda.md +++ b/src/Data/Int/Properties.lagda.md @@ -6,7 +6,8 @@ open import Data.Nat.Properties open import Data.Nat.Solver open import Data.Nat.Order open import Data.Int.Base -open import Data.Nat.Base +open import Data.Nat.Base hiding (Positive) +open import Data.Sum.Base ``` --> @@ -358,19 +359,14 @@ no further comments. lemma : ∀ x y z → z + y * suc z + x * suc (z + y * suc z) ≡ z + (y + x * suc y) * suc z lemma x y z = nat! + -- *ℤ-def : ∀ x y → x *ℤ y ≡ assign (sign× (sign x) (sign y)) (abs x * abs y) + *ℤ-associative : ∀ x y z → x *ℤ (y *ℤ z) ≡ (x *ℤ y) *ℤ z *ℤ-associative posz y z = refl - *ℤ-associative x posz z = - ap (λ e → assign (sign× (sign x) pos) e) (*-zeror (abs x)) ∙ sym (ap - (λ e → assign - (sign× (sign (assign (sign× (sign x) pos) e)) (sign z)) - (abs (assign (sign× (sign x) pos) e) * abs z)) (*-zeror (abs x))) - *ℤ-associative x y posz = - ap (λ e → - assign (sign× (sign x) (sign (assign (sign× (sign y) pos) e))) - (abs x * abs (assign (sign× (sign y) pos) e))) (*-zeror (abs y)) - ∙ ap₂ assign refl (*-zeror (abs x)) - ∙ sym (ap₂ assign refl (*-zeror (abs (assign (sign× (sign x) (sign y)) (abs x * abs y))))) + *ℤ-associative (pos x@(suc _)) posz z = ap (assign pos) (*-zeror x) ∙ sym (ap₂ _*ℤ_ (ap (assign pos) (*-zeror x)) refl) + *ℤ-associative (negsuc x) posz z = ap (assign neg) (*-zeror x) ∙ sym (ap₂ _*ℤ_ (ap (assign neg) (*-zeror x)) refl) + *ℤ-associative x (pos y) posz = ap₂ _*ℤ_ (λ i → x) (ap (assign pos) (*-zeror y)) ∙ *ℤ-zeror x ∙ sym (*ℤ-zeror (x *ℤ pos y)) + *ℤ-associative x (negsuc y) posz = ap₂ _*ℤ_ (λ i → x) (ap (assign neg) (*-zeror y)) ∙ *ℤ-zeror x ∙ sym (*ℤ-zeror (x *ℤ negsuc y)) *ℤ-associative (possuc x) (possuc y) (possuc z) = ap possuc (lemma x y z) *ℤ-associative (possuc x) (possuc y) (negsuc z) = ap negsuc (lemma x y z) *ℤ-associative (possuc x) (negsuc y) (possuc z) = ap negsuc (lemma x y z) @@ -404,7 +400,10 @@ no further comments. sign×-commutative neg neg = refl *ℤ-commutative : ∀ x y → x *ℤ y ≡ y *ℤ x - *ℤ-commutative x y = ap₂ assign (sign×-commutative _ _) (*-commutative (abs x) (abs y)) + *ℤ-commutative (pos x) (pos y) = ap (assign pos) (*-commutative x y) + *ℤ-commutative (pos x) (negsuc y) = ap (assign neg) (*-commutative x (suc y)) + *ℤ-commutative (negsuc x) (pos y) = ap (assign neg) (*-commutative (suc x) y) + *ℤ-commutative (negsuc x) (negsuc y) = ap (assign pos) (*-commutative (suc x) (suc y)) dot-is-mul : ∀ x y → (dotℤ x y) ≡ (x *ℤ y) dot-is-mul posz y = refl @@ -468,4 +467,68 @@ no further comments. *ℤ-injectiver posz x y k≠0 p = absurd (k≠0 refl) *ℤ-injectiver (possuc k) x y k≠0 p = *ℤ-injectiver-possuc k x y p *ℤ-injectiver (negsuc k) x y k≠0 p = *ℤ-injectiver-negsuc k x y p + + *ℤ-is-zero : ∀ x y → (x *ℤ y) ≡ 0 → (x ≡ 0) ⊎ (y ≡ 0) + *ℤ-is-zero posz (pos _) _ = inl refl + *ℤ-is-zero (negsuc _) posz _ = inr refl + *ℤ-is-zero posz (negsuc _) _ = inl refl + *ℤ-is-zero (possuc _) posz _ = inr refl + *ℤ-is-zero (possuc _) (possuc _) p = absurd (suc≠zero (pos-injective p)) + *ℤ-is-zero (possuc _) (negsuc _) p = absurd (negsuc≠pos p) + *ℤ-is-zero (negsuc _) (possuc _) p = absurd (negsuc≠pos p) + *ℤ-is-zero (negsuc _) (negsuc _) p = absurd (suc≠zero (pos-injective p)) + + abs-assign : ∀ s n → abs (assign s n) ≡ n + abs-assign pos zero = refl + abs-assign pos (suc n) = refl + abs-assign neg zero = refl + abs-assign neg (suc n) = refl + + abs-*ℤ : ∀ x y → abs (x *ℤ y) ≡ abs x * abs y + abs-*ℤ (pos x) (pos y) = abs-assign pos (x * y) + abs-*ℤ (pos x) (negsuc y) = abs-assign neg (x * suc y) + abs-*ℤ (negsuc x) (pos y) = abs-assign neg (y + x * y) + abs-*ℤ (negsuc x) (negsuc y) = refl + + sign-*ℤ-positive : ∀ x t → Positive t → sign (x *ℤ t) ≡ sign x + sign-*ℤ-positive (pos x) (possuc y) (pos _) = ap sign (assign-pos (x * suc y)) + sign-*ℤ-positive (negsuc x) (possuc y) (pos _) = refl + + assign-sign : ∀ x → assign (sign x) (abs x) ≡ x + assign-sign posz = refl + assign-sign (possuc _) = refl + assign-sign (negsuc _) = refl + + assign-pos-positive : ∀ x → Positive x → assign pos (abs x) ≡ x + assign-pos-positive x@(possuc _) (pos _) = refl + + divides-*ℤ : ∀ {n k m} → k * n ≡ abs m → pos k *ℤ assign (sign m) n ≡ m + divides-*ℤ {zero} {k} {pos x} p = ap (assign pos) (*-zeror k) ∙ ap Int.pos (sym (*-zeror k) ∙ p) + divides-*ℤ {suc n} {k} {posz} p = ap (assign pos) p + divides-*ℤ {suc n} {zero} {possuc x} p = ap pos p + divides-*ℤ {suc n} {suc k} {possuc x} p = ap pos p + divides-*ℤ {zero} {k} {negsuc x} p = absurd (zero≠suc (sym (*-zeror k) ∙ p)) + divides-*ℤ {suc n} {zero} {negsuc x} p = absurd (zero≠suc p) + divides-*ℤ {suc n} {suc k} {negsuc x} p = ap negsuc (suc-inj p) +``` + +## Positivity + +```agda +*ℤ-positive : ∀ {x y} → Positive x → Positive y → Positive (x *ℤ y) +*ℤ-positive (pos x) (pos y) = pos (y + x * suc y) + ++ℤ-positive : ∀ {x y} → Positive x → Positive y → Positive (x +ℤ y) ++ℤ-positive (pos x) (pos y) = pos (x + suc y) + +pos-positive : ∀ {x} → x ≠ 0 → Positive (pos x) +pos-positive {zero} 0≠0 = absurd (0≠0 refl) +pos-positive {suc n} _ = pos n + +positive→nonzero : ∀ {x} → Positive x → x ≠ 0 +positive→nonzero .{possuc x} (pos x) p = suc≠zero (pos-injective p) + +*ℤ-positive-cancel : ∀ {x y} → Positive x → Positive (x *ℤ y) → Positive y +*ℤ-positive-cancel {pos .(suc x)} {posz} (pos x) q = absurd (positive→nonzero q (ap (assign pos) (*-zeror x))) +*ℤ-positive-cancel {pos .(suc x)} {possuc y} (pos x) q = pos y ``` diff --git a/src/Data/List/Properties.lagda.md b/src/Data/List/Properties.lagda.md index 7354da83c..1e876b92c 100644 --- a/src/Data/List/Properties.lagda.md +++ b/src/Data/List/Properties.lagda.md @@ -92,8 +92,8 @@ We use this to prove that lists preserve h-levels for $n \ge 2$, i.e. if List-is-hlevel n ahl x y = Equiv→is-hlevel (suc n) Code≃Path Code-is-hlevel where Code-is-hlevel : {x y : List A} → is-hlevel (Code x y) (suc n) Code-is-hlevel {[]} {[]} = is-prop→is-hlevel-suc λ x y → refl - Code-is-hlevel {[]} {x ∷ y} = is-prop→is-hlevel-suc λ x → absurd (Lift.lower x) - Code-is-hlevel {x ∷ x₁} {[]} = is-prop→is-hlevel-suc λ x → absurd (Lift.lower x) + Code-is-hlevel {[]} {x ∷ y} = is-prop→is-hlevel-suc λ x → absurd (lower x) + Code-is-hlevel {x ∷ x₁} {[]} = is-prop→is-hlevel-suc λ x → absurd (lower x) Code-is-hlevel {x ∷ x₁} {x₂ ∷ y} = ×-is-hlevel (suc n) (ahl _ _) Code-is-hlevel instance diff --git a/src/Data/Nat/Base.lagda.md b/src/Data/Nat/Base.lagda.md index ae1de2e0f..55df78a59 100644 --- a/src/Data/Nat/Base.lagda.md +++ b/src/Data/Nat/Base.lagda.md @@ -219,6 +219,10 @@ instance {-# INCOHERENT x≤x x≤sucy #-} +factorial : Nat → Nat +factorial zero = 1 +factorial (suc n) = suc n * factorial n + Positive : Nat → Type Positive n = 1 ≤ n ``` diff --git a/src/Data/Nat/Divisible.lagda.md b/src/Data/Nat/Divisible.lagda.md index 266ae70b5..203d9f929 100644 --- a/src/Data/Nat/Divisible.lagda.md +++ b/src/Data/Nat/Divisible.lagda.md @@ -150,6 +150,9 @@ expect a number to divide its multiples. Fortunately, this is the case: |-*l-pres : ∀ {n a b} → n ∣ b → n ∣ a * b |-*l-pres {n} {a} {b} p1 with (q , α) ← ∣→fibre p1 = fibre→∣ (a * q , *-associative a q n ∙ ap (a *_) α) + +∣-*-cancelr : ∀ {n a b} .⦃ _ : Positive n ⦄ → a * n ∣ b * n → a ∣ b +∣-*-cancelr {n} {a} {b} p1 with (q , α) ← ∣→fibre p1 = fibre→∣ (q , *-injr n (q * a) b (*-associative q a n ∙ α)) ``` If two numbers are multiples of $k$, then so is their sum. diff --git a/src/Data/Nat/Divisible/GCD.lagda.md b/src/Data/Nat/Divisible/GCD.lagda.md index 5eb910793..8d3ca1bbc 100644 --- a/src/Data/Nat/Divisible/GCD.lagda.md +++ b/src/Data/Nat/Divisible/GCD.lagda.md @@ -193,6 +193,24 @@ is-gcd-graphs-gcd {m = m} {n} {d} = prop-ext! (λ p → subst (is-gcd m n) p (Euclid.euclid m n .snd)) ``` + + ## Euclid's lemma ```agda diff --git a/src/Data/Nat/Prime.lagda.md b/src/Data/Nat/Prime.lagda.md index 68f1d668a..4c6ff0f25 100644 --- a/src/Data/Nat/Prime.lagda.md +++ b/src/Data/Nat/Prime.lagda.md @@ -256,18 +256,6 @@ factorisation-worker n@(suc (suc m)) ind with is-prime-or-composite n (s≤s (s ; is-primes = composite .p-prime ∷ primes } -factorial : Nat → Nat -factorial zero = 1 -factorial (suc n) = suc n * factorial n - -factorial-positive : ∀ n → Positive (factorial n) -factorial-positive zero = s≤s 0≤x -factorial-positive (suc n) = *-preserves-≤ 1 (suc n) 1 (factorial n) (s≤s 0≤x) (factorial-positive n) - -≤-factorial : ∀ n → n ≤ factorial n -≤-factorial zero = 0≤x -≤-factorial (suc n) = subst (_≤ factorial (suc n)) (*-oner (suc n)) (*-preserves-≤ (suc n) (suc n) 1 (factorial n) ≤-refl (factorial-positive n)) - ∣-factorial : ∀ n {m} → m < n → suc m ∣ factorial n ∣-factorial (suc n) {m} m≤n with suc m ≡? suc n ... | yes m=n = subst (_∣ factorial (suc n)) (sym m=n) (∣-*l {suc n} {factorial n}) diff --git a/src/Data/Nat/Properties.lagda.md b/src/Data/Nat/Properties.lagda.md index a4b941330..2cb328a9d 100644 --- a/src/Data/Nat/Properties.lagda.md +++ b/src/Data/Nat/Properties.lagda.md @@ -6,6 +6,7 @@ open import 1Lab.Type open import Data.Nat.Order open import Data.Dec.Base open import Data.Nat.Base +open import Data.Sum.Base open import Data.Bool ``` --> @@ -134,6 +135,21 @@ numbers]. Since they're mostly simple inductive arguments written in *-is-oner : ∀ x n → x * n ≡ 1 → n ≡ 1 *-is-oner x n p = *-is-onel n x (*-commutative n x ∙ p) + +*-is-zero : ∀ x y → x * y ≡ 0 → (x ≡ 0) ⊎ (y ≡ 0) +*-is-zero zero y p = inl refl +*-is-zero (suc x) zero p = inr refl +*-is-zero (suc x) (suc y) p = absurd (suc≠zero p) + +*-is-zerol : ∀ x y ⦃ _ : Positive y ⦄ → x * y ≡ 0 → x ≡ 0 +*-is-zerol x (suc y) p with *-is-zero x (suc y) p +... | inl p = p +... | inr q = absurd (suc≠zero q) + +*-is-zeror : ∀ x y ⦃ _ : Positive x ⦄ → x * y ≡ 0 → y ≡ 0 +*-is-zeror (suc x) y p with *-is-zero (suc x) y p +... | inl p = absurd (suc≠zero p) +... | inr q = q ``` ## Exponentiation @@ -252,9 +268,14 @@ difference→≤ {suc x} {suc z} (suc y) p = s≤s (difference→≤ (suc y) (su nonzero→positive : ∀ {x} → x ≠ 0 → 0 < x nonzero→positive {zero} p = absurd (p refl) nonzero→positive {suc x} p = s≤s 0≤x + +*-cancel-≤r : ∀ x {y z} .⦃ _ : Positive x ⦄ → (y * x) ≤ (z * x) → y ≤ z +*-cancel-≤r (suc x) {zero} {z} p = 0≤x +*-cancel-≤r (suc x) {suc y} {suc z} (s≤s p) = s≤s + (*-cancel-≤r (suc x) {y} {z} (+-reflects-≤l (y * suc x) (z * suc x) x p)) ``` -### Monus +## Monus ```agda monus-zero : ∀ a → 0 - a ≡ 0 @@ -294,7 +315,7 @@ monus-swapr : ∀ x y z → x + y ≡ z → x ≡ z - y monus-swapr x y z p = sym (monus-cancelr x 0 y) ∙ ap (_- y) p ``` -### Maximum +## Maximum ```agda max-assoc : (x y z : Nat) → max x (max y z) ≡ max (max x y) z @@ -334,7 +355,7 @@ max-zeror zero = refl max-zeror (suc x) = refl ``` -### Minimum +## Minimum ```agda min-assoc : (x y z : Nat) → min x (min y z) ≡ min (min x y) z @@ -371,3 +392,15 @@ min-zeror : (x : Nat) → min x 0 ≡ 0 min-zeror zero = refl min-zeror (suc x) = refl ``` + +## The factorial function + +```agda +factorial-positive : ∀ n → Positive (factorial n) +factorial-positive zero = s≤s 0≤x +factorial-positive (suc n) = *-preserves-≤ 1 (suc n) 1 (factorial n) (s≤s 0≤x) (factorial-positive n) + +≤-factorial : ∀ n → n ≤ factorial n +≤-factorial zero = 0≤x +≤-factorial (suc n) = subst (_≤ factorial (suc n)) (*-oner (suc n)) (*-preserves-≤ (suc n) (suc n) 1 (factorial n) ≤-refl (factorial-positive n)) +``` diff --git a/src/Data/Rational/Base.lagda.md b/src/Data/Rational/Base.lagda.md new file mode 100644 index 000000000..40bc1a12f --- /dev/null +++ b/src/Data/Rational/Base.lagda.md @@ -0,0 +1,626 @@ + + +```agda +module Data.Rational.Base where +``` + +# Rational numbers {defines="rational-numbers"} + +The ring of **rational numbers**, $\bQ$, is the +[[localisation|localisation of a ring]] of the ring of [[integers]] +$\bZ$, inverting every positive element. We've already done most of the +work while implementing localisations: + +```agda +PositiveΩ : Int → Ω +PositiveΩ x .∣_∣ = Positive x +PositiveΩ x .is-tr = hlevel 1 + +private + module L = Loc ℤ-comm PositiveΩ record { has-1 = pos 0 ; has-* = *ℤ-positive } +``` + + + +Strictly speaking, we are done: we could simply define $\bQ$ to be the +ring we just constructed. However, for the sake of implementation +hiding, we wrap it as a distinct type constructor. This lets consumers +of the type $\bQ$ forget that it's implemented as a localisation. + +```agda +data ℚ : Type where + inc : ⌞ L.S⁻¹R ⌟ → ℚ + +toℚ : Fraction → ℚ +toℚ x = inc (inc x) + +_+ℚ_ : ℚ → ℚ → ℚ +_+ℚ_ (inc x) (inc y) = inc (x L.+ₗ y) + +_*ℚ_ : ℚ → ℚ → ℚ +_*ℚ_ (inc x) (inc y) = inc (x L.*ₗ y) + +-ℚ_ : ℚ → ℚ +-ℚ_ (inc x) = inc (L.-ₗ x) +``` + + + +However, clients of this module *will* need the fact that $\bQ$ is a +quotient of the type of integer fractions. Therefore, we expose an +elimination principle, saying that to show a [[proposition]] everywhere +over $\bQ$, it suffices to do so at the fractions. + +```agda +ℚ-elim-prop + : ∀ {ℓ} {P : ℚ → Type ℓ} (pprop : ∀ x → is-prop (P x)) + → (f : ∀ x → P (toℚ x)) + → ∀ x → P x +ℚ-elim-prop pprop f (inc (inc x)) = f x +ℚ-elim-prop pprop f (inc (glue r@(x , y , _) i)) = is-prop→pathp (λ i → pprop (inc (glue r i))) (f x) (f y) i +ℚ-elim-prop pprop f (inc (squash x y p q i j)) = + is-prop→squarep + (λ i j → pprop (inc (squash x y p q i j))) + (λ i → go (inc x)) (λ i → go (inc (p i))) (λ i → go (inc (q i))) (λ i → go (inc y)) + i j + where go = ℚ-elim-prop pprop f +``` + + + +Next, we show that sameness of fractions implies identity in $\bQ$, and +the converse is true as well: + +```agda +abstract + quotℚ : ∀ {x y} → x ≈ y → toℚ x ≡ toℚ y + quotℚ p = ap ℚ.inc (quot p) + + unquotℚ : ∀ {x y} → toℚ x ≡ toℚ y → x ≈ y + unquotℚ p = ≈.effective (ap unℚ p) +``` + +Finally, we want to show that the type of rational numbers is discrete. +To do this, we have to show that sameness of integer fractions is +decidable, i.e. that, given fractions $x/s$ and $y/t$, we can decide +whether there exists a $u \ne 0$ with $uxt = uys$. This is not automatic +since $u$ can range over all integers! However, recall that this is +equivalent to $u(xt - ys) = 0$. Since we know that $\bZ$ has no +zerodivisors, if this is true, then either $u = 0$ or $xt - ys = 0$; by +assumption, it can not be $u$. But if $xt - ys = 0$, then $xt = ys$! +Conversely, if $xt = ys$, then we can take $u = 1$. Therefore, checking +sameness of fractions boils down to checking whether they +"cross-multiply" to the same thing. + +```agda +from-same-rational : {x y : Fraction} → x ≈ y → x .↑ *ℤ y .↓ ≡ y .↑ *ℤ x .↓ +from-same-rational {x / s [ s≠0 ]} {y / t [ t≠0 ]} p = case L.≈→≈' p of λ where + u@(possuc u') (pos u') p → case *ℤ-is-zero u _ p of λ where + (inl u=0) → absurd (suc≠zero (pos-injective u=0)) + (inr xt-ys=0) → ℤ.zero-diff xt-ys=0 + +to-same-rational : {x y : Fraction} → x .↑ *ℤ y .↓ ≡ y .↑ *ℤ x .↓ → x ≈ y +to-same-rational {x / s [ s≠0 ]} {y / t [ t≠0 ]} p = L.inc 1 (pos 0) (recover (sym (*ℤ-associative 1 x t) ·· ap (1 *ℤ_) p ·· *ℤ-associative 1 y s)) + +Dec-same-rational : (x y : Fraction) → Dec (x ≈ y) +Dec-same-rational f@(x / s [ _ ]) f'@(y / t [ _ ]) with x *ℤ t ≡? y *ℤ s +... | yes p = yes (to-same-rational p) +... | no xt≠ys = no λ p → xt≠ys (from-same-rational p) +``` + + + +
+ +There are a number of other properties of $\bZ{\loc{\ne 0}}$ that we can +export as properties of $\bQ$; however, these are all trivial as above, +so we do not remark on them any further. + + +```agda +_-ℚ_ : ℚ → ℚ → ℚ +x -ℚ y = x +ℚ (-ℚ y) + +infixl 8 _+ℚ_ _-ℚ_ +infixl 9 _*ℚ_ +infix 10 -ℚ_ + +_/_ : (x y : Int) ⦃ _ : Positive y ⦄ → ℚ +_/_ x y ⦃ p ⦄ = toℚ (x / y [ p ]) + +infix 11 _/_ + +{-# DISPLAY ℚ.inc (Coeq.inc (_/_[_] x y p)) = x / y #-} + +_/1 : Int → ℚ +x /1 = x / 1 + +instance + H-Level-ℚ : ∀ {n} → H-Level ℚ (2 + n) + H-Level-ℚ = basic-instance 2 (Discrete→is-set auto) + + Number-ℚ : Number ℚ + Number-ℚ .Number.Constraint _ = ⊤ + Number-ℚ .Number.fromNat x = pos x /1 + + Negative-ℚ : Negative ℚ + Negative-ℚ .Negative.Constraint _ = ⊤ + Negative-ℚ .Negative.fromNeg 0 = 0 + Negative-ℚ .Negative.fromNeg (suc x) = negsuc x /1 + + Inductive-ℚ + : ∀ {ℓ ℓm} {P : ℚ → Type ℓ} + → ⦃ _ : Inductive ((x : Fraction) → P (toℚ x)) ℓm ⦄ + → ⦃ _ : ∀ {x} → H-Level (P x) 1 ⦄ + → Inductive (∀ x → P x) ℓm + Inductive-ℚ ⦃ r ⦄ .Inductive.methods = r .Inductive.methods + Inductive-ℚ ⦃ r ⦄ .Inductive.from f = ℚ-elim-prop (λ x → hlevel 1) (r .Inductive.from f) + +abstract + +ℚ-idl : ∀ x → 0 +ℚ x ≡ x + +ℚ-idl (inc x) = ap inc (L.+ₗ-idl x) + + +ℚ-idr : ∀ x → x +ℚ 0 ≡ x + +ℚ-idr (inc x) = ap ℚ.inc (CRing.+-idr L.S⁻¹R) + + +ℚ-associative : ∀ x y z → x +ℚ (y +ℚ z) ≡ (x +ℚ y) +ℚ z + +ℚ-associative (inc x) (inc y) (inc z) = ap inc (L.+ₗ-assoc x y z) + + +ℚ-commutative : ∀ x y → x +ℚ y ≡ y +ℚ x + +ℚ-commutative (inc x) (inc y) = ap inc (L.+ₗ-comm x y) + + *ℚ-idl : ∀ x → 1 *ℚ x ≡ x + *ℚ-idl (inc x) = ap inc (L.*ₗ-idl x) + + *ℚ-idr : ∀ x → x *ℚ 1 ≡ x + *ℚ-idr (inc x) = ap ℚ.inc (CRing.*-idr L.S⁻¹R) + + *ℚ-associative : ∀ x y z → x *ℚ (y *ℚ z) ≡ (x *ℚ y) *ℚ z + *ℚ-associative (inc x) (inc y) (inc z) = ap inc (L.*ₗ-assoc x y z) + + *ℚ-commutative : ∀ x y → x *ℚ y ≡ y *ℚ x + *ℚ-commutative (inc x) (inc y) = ap inc (L.*ₗ-comm x y) + + *ℚ-zerol : ∀ x → 0 *ℚ x ≡ 0 + *ℚ-zerol (inc x) = ap ℚ.inc (CRing.*-zerol L.S⁻¹R {x}) + + *ℚ-zeror : ∀ x → x *ℚ 0 ≡ 0 + *ℚ-zeror (inc x) = ap ℚ.inc (CRing.*-zeror L.S⁻¹R {x}) + + *ℚ-distribl : ∀ x y z → x *ℚ (y +ℚ z) ≡ x *ℚ y +ℚ x *ℚ z + *ℚ-distribl (inc x) (inc y) (inc z) = ap ℚ.inc (L.*ₗ-distribl x y z) + ++ℚ-monoid : is-monoid 0 _+ℚ_ ++ℚ-monoid = record { has-is-semigroup = record { has-is-magma = record { has-is-set = hlevel 2 } ; associative = λ {x} {y} {z} → +ℚ-associative x y z } ; idl = +ℚ-idl _ ; idr = +ℚ-idr _ } + +*ℚ-monoid : is-monoid 1 _*ℚ_ +*ℚ-monoid = record { has-is-semigroup = record { has-is-magma = record { has-is-set = hlevel 2 } ; associative = λ {x} {y} {z} → *ℚ-associative x y z } ; idl = *ℚ-idl _ ; idr = *ℚ-idr _ } +``` + +
+ +## As a field + +An important property of the ring $\bQ$ is that any nonzero rational +number is invertible. Since inverses are unique when they exist --- the +type of inverses is a proposition --- it suffices to show this at the +more concrete level of integer fractions. + +```agda +inverseℚ : (x : ℚ) → x ≠ 0 → Σ[ y ∈ ℚ ] (x *ℚ y ≡ 1) +inverseℚ = ℚ-elim-prop is-p go where + abstract + is-p : (x : ℚ) → is-prop (x ≠ 0 → Σ[ y ∈ ℚ ] (x *ℚ y ≡ 1)) + is-p x = Π-is-hlevel 1 λ _ (y , p) (z , q) → Σ-prop-path! (monoid-inverse-unique *ℚ-monoid x y z (*ℚ-commutative y x ∙ p) q) + + rem₁ : ∀ x y → 1 *ℤ (x *ℤ y) *ℤ 1 ≡ 1 *ℤ (y *ℤ x) + rem₁ x y = ap (_*ℤ 1) (*ℤ-onel (x *ℤ y)) + ∙ *ℤ-oner (x *ℤ y) ∙ *ℤ-commutative x y ∙ sym (*ℤ-onel (y *ℤ x)) +``` + +This leaves us with three cases: either the fraction $x/y$ had a +denominator of zero, contradicting our assumption, or its numerator is +also nonzero (either positive or negative), so that we can form the +fraction $y/x$. It's then routine to show that $(x/y)(y/x) = 1$ holds in +$\bQ$. + +```agda + go : (x : Fraction) → toℚ x ≠ 0 → Σ[ y ∈ ℚ ] (toℚ x *ℚ y ≡ 1) + go (posz / y [ _ ]) nz = absurd (nz (quotℚ (L.inc 1 decide! refl))) + go (x@(possuc x') / y [ _ ]) nz = y / x , quotℚ (L.inc 1 decide! (rem₁ x y)) + go (x@(negsuc x') / y [ p ]) nz with y | p + ... | possuc y | _ = negsuc y / possuc x' , quotℚ (L.inc 1 decide! (rem₁ (negsuc x') (negsuc y))) + ... | negsuc y | _ = possuc y / possuc x' , quotℚ (L.inc 1 decide! (rem₁ x (possuc y))) +``` + +## Reduced fractions + +We now show that the quotient defining $\bQ$ is [[*split*|split +quotient]]: integer fractions have a well-defined notion of *normal +form*, and similarity of fractions is equivalent to equality under +normalisation. The procedure we formalise is the familiar one, sending a +fraction $x/y$ to + +$$ +\frac{x \ndiv \gcd(x, y)}{y \ndiv \gcd(x, y)} +$$. + +What remains is proving that this is actually a normalisation procedure, +which takes up the remainder of this module. + +```agda +reduce-fraction : Fraction → Fraction +reduce-fraction (x / y [ p ]) = reduced module reduce where + gcd[x,y] : GCD (abs x) (abs y) + gcd[x,y] = Euclid.euclid (abs x) (abs y) +``` + + + +Our first obligation, to form the reduced fraction at all, is showing +that the denominator is non-zero. This is pretty easy: we know that $y$ +is nonzero, so if $y \ndiv \gcd(x,y)$ were zero, we would have +$$y = \gcd(x, y) * (y \ndiv \gcd(x,y)) = \gcd(x,y) * 0 = 0$$, +which is a contradiction. A similar argument shows that $\gcd(x,y)$ is +itself nonzero, a fact we'll need later. + +```agda + rem₁ : y/g ≠ 0 + rem₁ y/g=0 with y/g | y/g=0 | y/g*g=y + ... | zero | y/g=0 | q = positive→nonzero p (abs-positive y (sym q)) + ... | suc n | y/g=0 | q = absurd (suc≠zero y/g=0) + + rem₂ : g ≠ 0 + rem₂ g=0 = positive→nonzero p (abs-positive y (sym (sym (*-zeror y/g) ∙ ap (y/g *_) (sym g=0) ∙ y/g*g=y))) +``` + +Finally, we must quickly mention the issue of signs. Since `gcd`{.Agda} +produces a natural number, we have to multiply it by the sign $\sgn(x)$ +of $x$, to make sure signs are preserved. Note that the denominator must +be positive. + +```agda + reduced : Fraction + reduced = assign (sign x) x/g / pos y/g [ pos-positive rem₁ ] +``` + +Finally, we can calculate that these fractions are similar. + +```agda + related : (x / y [ p ]) ≈ reduced + related = L.inc (pos g) (pos-positive rem₂) $ + pos g *ℤ x *ℤ pos y/g ≡⟨ solve 3 (λ x y z → x :* y :* z ≔ (z :* x) :* y) (pos g) x (pos y/g) refl ⟩ + (pos y/g *ℤ pos g) *ℤ x ≡⟨ ap (_*ℤ x) (ap (assign pos) y/g*g=y) ⟩ + assign pos (abs y) *ℤ x ≡⟨ ap₂ _*ℤ_ (assign-pos-positive y p) (sym (divides-*ℤ {n = x/g} {g} {x} (*-commutative g x/g ∙ x/g*g=x))) ⟩ + y *ℤ (pos g *ℤ assign (sign x) x/g) ≡⟨ solve 3 (λ y g x → y :* (g :* x) ≔ g :* x :* y) y (pos g) (assign (sign x) x/g) refl ⟩ + pos g ℤ.* assign (sign x) x/g ℤ.* y ∎ +``` + + + +We'll now show that `reduce-fraction`{.Agda} respects similarity of +fractions. We show this by an intermediate lemma, which says that, given +a non-zero $s$ and a fraction $x/y$, we have +$$ +\frac{xs \ndiv \gcd(xs, ys)}{ys \ndiv \gcd(xs, ys)} = \frac{x \ndiv \gcd(x, y)}{y \ndiv \gcd(x, y)} +$$, +as integer fractions (rather than rational numbers). + +```agda +reduce-*r + : ∀ x y s (p : Positive y) (q : Positive s) + → reduce-fraction ((x *ℤ s) / (y *ℤ s) [ *ℤ-positive p q ]) + ≡ reduce-fraction (x / y [ p ]) +reduce-*r x y s p q = Fraction-path (ap₂ assign sgn num) (ap Int.pos denom) where +``` + + + +The first observation is that $\gcd(xs, ys) = \gcd(x,y)s$, even when +absolute values are involved.^[Keep in mind that the `gcd`{.Agda} +function is defined over the naturals, and we're extending it to +integers by $\gcd(x,y) = \gcd(\abs{x}, \abs{y})$.] + +```agda + p1 : gcdℤ (x *ℤ s) (y *ℤ s) ≡ gcdℤ x y * abs s + p1 = ap₂ gcd (abs-*ℤ x s) (abs-*ℤ y s) ∙ gcd-factor (abs x) (abs y) (abs s) +``` + +This implies that $$(xs \ndiv \gcd(xs, ys)) \gcd(x,y) s = x s$$, and, +because multiplication by $s$ is injective, this in turn implies that +$(xs \ndiv \gcd(xs, ys))\gcd(x, y)$ is also $x$. We conclude $(xs \ndiv +\gcd(xs, ys)) = (x \ndiv \gcd(x, y))$, since both sides multiply with +$\gcd(x, y)$ to $x$, and this multiplication is *also* injective. +Therefore, our numerators have the same absolute value. + +```agda + num' : xs/gcd * gcdℤ x y ≡ abs x + num' = *-injr (abs s) (xs/gcd * m.g) (abs x) $ + xs/gcd * gcdℤ x y * abs s ≡⟨ *-associative xs/gcd m.g (abs s) ⟩ + xs/gcd * (gcdℤ x y * abs s) ≡˘⟨ ap (xs/gcd *_) p1 ⟩ + xs/gcd * gcdℤ (x *ℤ s) (y *ℤ s) ≡⟨ n.x/g*g=x ⟩ + abs (x *ℤ s) ≡⟨ abs-*ℤ x s ⟩ + abs x * abs s ∎ + + num : xs/gcd ≡ x/gcd + num = *-injr (gcdℤ x y) xs/gcd x/gcd (num' ∙ sym m.x/g*g=x) +``` + +We must still show that the resulting numerators have the same sign. +Fortunately, this boils down to a bit of algebra, plus the +hyper-specific fact that $\sgn(ab^2) = \sgn(a)$, whenever $b$ is +nonzero.^[Here, we're applying this lemma with $a = xy$ and $b = s$, and +we have assumed $s$ nonzero. The $\sgn$ function is not fun.] + +```agda + sgn : sign (x *ℤ s) ≡ sign x + sgn = sign-*ℤ-positive x s q +``` + +A symmetric calculation shows that the denominators also have the same +absolute value, and, since they're both positive in the resulting +fraction, we're done. + +```agda + denom' : ys/gcd * gcdℤ x y ≡ abs y + denom' = *-injr (abs s) (ys/gcd * m.g) (abs y) (*-associative ys/gcd m.g (abs s) ∙ sym (ap (ys/gcd *_) p1) ∙ n.y/g*g=y ∙ abs-*ℤ y s) + + denom : ys/gcd ≡ y/gcd + denom = *-injr (gcdℤ x y) ys/gcd y/gcd (denom' ∙ sym m.y/g*g=y) +``` + +We can use this to show that `reduce-fraction`{.Agda} sends similar +fractions to *equal* normal forms: If $x/s \approx y/t$, we have $xt = +ys$, and we can calculate +$$ +\frac{x \ndiv \gcd(x, s)}{s \ndiv \gcd(x, s)} += \frac{xt \ndiv \gcd(xt, st)}{st \ndiv \gcd(xt, st)} += \frac{ys \ndiv \gcd(ys, ts)}{ts \ndiv \gcd(ys, ts)} += \frac{y \ndiv \gcd(y, t)}{t \ndiv \gcd(y, t)} +$$ +using the previous lemma. Thus, by the general logic of [[split +quotients]], we conclude that $\bQ$ is equivalent to the set of reduced +integer fractions. + +```agda +reduce-resp : (x y : Fraction) → x ≈ y → reduce-fraction x ≡ reduce-fraction y +reduce-resp f@(x / s [ s≠0 ]) f'@(y / t [ t≠0 ]) p = + reduce-fraction (x / s [ s≠0 ]) ≡⟨ sym (reduce-*r x s t s≠0 t≠0) ⟩ + reduce-fraction ((x *ℤ t) / (s *ℤ t) [ _ ]) ≡⟨ ap reduce-fraction (Fraction-path {x = _ / _ [ *ℤ-positive s≠0 t≠0 ]} {_ / _ [ *ℤ-positive t≠0 s≠0 ]} (from-same-rational p) (*ℤ-commutative s t)) ⟩ + reduce-fraction ((y *ℤ s) / (t *ℤ s) [ _ ]) ≡⟨ reduce-*r y t s t≠0 s≠0 ⟩ + reduce-fraction (y / t [ t≠0 ]) ∎ + +integer-frac-splits : is-split-congruence L.Fraction-congruence +integer-frac-splits = record + { has-is-set = hlevel 2 + ; normalise = reduce-fraction + ; represents = elim! reduce.related + ; respects = reduce-resp + } +``` + + diff --git a/src/Data/Rational/Order.lagda.md b/src/Data/Rational/Order.lagda.md new file mode 100644 index 000000000..120af8de6 --- /dev/null +++ b/src/Data/Rational/Order.lagda.md @@ -0,0 +1,239 @@ + + +```agda +module Data.Rational.Order where +``` + + + +# Ordering on rationals + +The field $\bQ$ of [[rational numbers]] inherits a [[partial order]] +from the ring of integers $\bZ$, defining the relation +$$ +\frac{x}{s} \le \frac{y}{t} +$$ +to be $xt \le ys$. As usual, we define this first at the level of +fractions, then prove that it extends to the quotient $\bQ$. + +```agda +_≤ᶠ_ : Fraction → Fraction → Type +(x / s [ _ ]) ≤ᶠ (y / t [ _ ]) = (x *ℤ t) ℤ.≤ (y *ℤ s) +``` + +The easiest way to show this is transitivity at the level of fractions, +since it will follow directly from the definition that $q \approx q'$ +implies $q \le q'$. We can then prove that $q \le q'$ is invariant under +$\approx$, on either side, by appealing to transitivity. The proof is +not complicated, just annoying: + +```agda +≤ᶠ-refl' : ∀ {x y} → x ≈ y → x ≤ᶠ y +≤ᶠ-refl' {x@record{}} {y@record{}} p = ℤ.≤-refl' (from-same-rational p) + +≤ᶠ-trans : ∀ x y z → x ≤ᶠ y → y ≤ᶠ z → x ≤ᶠ z +≤ᶠ-trans (x / s [ pos s' ]) (y / t [ pos t' ]) (z / u [ pos u' ]) α β = + (ℤ.*ℤ-cancel-≤r {t} $ + x *ℤ u *ℤ t =⟨ solve 3 (λ x u t → x :* u :* t ≔ x :* t :* u) x u t refl ⟩ + x *ℤ t *ℤ u ≤⟨ ℤ.*ℤ-preserves-≤r u α ⟩ + y *ℤ s *ℤ u =⟨ solve 3 (λ y s u → y :* s :* u ≔ y :* u :* s) y s u refl ⟩ + y *ℤ u *ℤ s ≤⟨ ℤ.*ℤ-preserves-≤r s β ⟩ + z *ℤ t *ℤ s =⟨ solve 3 (λ z t s → z :* t :* s ≔ z :* s :* t) z t s refl ⟩ + z *ℤ s *ℤ t ≤∎) +``` + +
+ +We can then lift this to a family of [[propositions]] over $\bQ \times +\bQ$. The strategy for showing it respects the quotient is as outlined +above, so we'll leave this in a `
`{.html} block. +
+ +```agda +private + leℚ : ℚ → ℚ → Prop lzero + leℚ (inc x) (inc y) = + Coeq-rec₂ (hlevel 2) work + (λ a (x , y , r) → r2 x y a r) + (λ a (x , y , r) → r1 a x y r) x y + where + work : Fraction → Fraction → Prop lzero + ∣ work x y ∣ = x ≤ᶠ y + work record{} record{} .is-tr = hlevel 1 + + r1 : ∀ u x y → x ≈ y → work u x ≡ work u y + r1 u@record{} x@record{} y@record{} p' = n-ua (prop-ext! + (λ α → ≤ᶠ-trans u x y α (≤ᶠ-refl' p')) + (λ α → ≤ᶠ-trans u y x α (≤ᶠ-refl' (≈.symᶜ p')))) + + r2 : ∀ x y u → x ≈ y → work x u ≡ work y u + r2 u@record{} x@record{} y@record{} p' = n-ua (prop-ext! + (λ α → ≤ᶠ-trans x u y (≤ᶠ-refl' (≈.symᶜ p')) α) + (λ α → ≤ᶠ-trans u x y (≤ᶠ-refl' p') α)) +``` + +
+ +We define the type family `_≤_`{.Agda} as a record type, wrapping +`orderℚ`{.Agda}, to help out type inference: `orderℚ`{.Agda} is a pretty +nasty definition, whereas an application of a record name is always a +normal form. + +```agda +record _≤_ (x y : ℚ) : Type where + constructor inc + field + lower : ⌞ leℚ x y ⌟ +``` + + + +By lifting our proofs about `_≤ᶠ_`{.Agda} through this record type, we +can prove that the ordering on the rational numbers is decidable, +reflexive, transitive, and antisymmetric. + +```agda +≤-dec : ∀ x y → Dec (x ≤ y) +≤-dec = elim! go where + go : ∀ x y → Dec (toℚ x ≤ toℚ y) + go x@record{} y@record{} with holds? ((↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x)) + ... | yes p = yes (inc p) + ... | no ¬p = no (¬p ∘ lower) + +instance + Dec-≤ : ∀ {x y} → Dec (x ≤ y) + Dec-≤ {x} {y} = ≤-dec x y + +abstract + toℚ≤ : ∀ {x y} → (↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x) → toℚ x ≤ toℚ y + toℚ≤ {record{}} {record{}} p = inc p + + ≤-refl : ∀ {x} → x ≤ x + ≤-refl {x} = work x where + work : ∀ x → x ≤ x + work = elim! λ f → toℚ≤ ℤ.≤-refl + + ≤-trans : ∀ {x y z} → x ≤ y → y ≤ z → x ≤ z + ≤-trans {x} {y} {z} = work x y z where + work : ∀ x y z → x ≤ y → y ≤ z → x ≤ z + work = elim! λ x y z p q → inc (≤ᶠ-trans x y z (p .lower) (q .lower)) + + ≤-antisym : ∀ {x y} → x ≤ y → y ≤ x → x ≡ y + ≤-antisym {x} {y} = work x y where + work : ∀ x y → x ≤ y → y ≤ x → x ≡ y + work = elim! λ where + x@record{} y@record{} p q → quotℚ (to-same-rational (ℤ.≤-antisym (p .lower) (q .lower))) +``` + +## Algebraic properties + +We can also show that the ordering on $\bQ$ behaves well with respect to +its algebraic structure. For example, the addition function is monotone +in both arguments, and division by a positive integer preserves the +order. + +```agda + common-denominator-≤ : ∀ {x y d} (p : ℤ.Positive d) → x ℤ.≤ y → toℚ (x / d [ p ]) ≤ toℚ (y / d [ p ]) + common-denominator-≤ {d = d} (pos _) p = inc (ℤ.*ℤ-preserves-≤r d p) + + +ℚ-preserves-≤ : ∀ x y a b → x ≤ y → a ≤ b → (x +ℚ a) ≤ (y +ℚ b) + +ℚ-preserves-≤ = ℚ-elim-propⁿ 4 λ d d≠0 x y a b (inc p) (inc q) → + common-denominator-≤ (ℤ.*ℤ-positive d≠0 d≠0) $ + x *ℤ d +ℤ a *ℤ d ≤⟨ ℤ.+ℤ-preserves-≤r (a *ℤ d) _ _ p ⟩ + y *ℤ d +ℤ a *ℤ d ≤⟨ ℤ.+ℤ-preserves-≤l (y *ℤ d) _ _ q ⟩ + y *ℤ d +ℤ b *ℤ d ≤∎ + + *ℚ-nonnegative : ∀ x y → 0 ≤ x → 0 ≤ y → 0 ≤ (x *ℚ y) + *ℚ-nonnegative = ℚ-elim-propⁿ 2 λ d d≠0 x y (inc p) (inc q) → inc (ℤ.≤-trans + (ℤ.*ℤ-nonnegative + (ℤ.≤-trans p (ℤ.≤-refl' (ℤ.*ℤ-oner x))) + (ℤ.≤-trans q (ℤ.≤-refl' (ℤ.*ℤ-oner y)))) + (ℤ.≤-refl' (sym (ℤ.*ℤ-oner (x *ℤ y))))) +``` + +## Positivity + +We can also lift the notion of positivity from the integers to the +rational numbers. A fraction is positive if its numerator is positive. + +```agda +private + positiveℚ : ℚ → Prop lzero + positiveℚ (inc x) = Coeq-rec (λ f → el! (ℤ.Positive (↑ f))) (λ (x , y , p) → r x y p) x where + r : ∀ x y → x ≈ y → el! (ℤ.Positive (↑ x)) ≡ el! (ℤ.Positive (↑ y)) + r (x / s [ ps ]) (y / t [ pt ]) r with r ← from-same-rational r = n-ua (prop-ext! + (λ px → ℤ.*ℤ-positive-cancel ps (subst ℤ.Positive (r ∙ ℤ.*ℤ-commutative y s) (ℤ.*ℤ-positive px pt))) + λ py → ℤ.*ℤ-positive-cancel pt (subst ℤ.Positive (sym r ∙ ℤ.*ℤ-commutative x t) (ℤ.*ℤ-positive py ps))) + +record Positive (q : ℚ) : Type where + constructor inc + field + lower : ⌞ positiveℚ q ⌟ +``` + + + +This has the expected interaction with the ordering and the algebraic +operations. + +```agda +nonnegative-nonzero→positive : ∀ {x} → 0 ≤ x → x ≠ 0 → Positive x +nonnegative-nonzero→positive = work _ where + work : ∀ x → 0 ≤ x → x ≠ 0 → Positive x + work = ℚ-elim-propⁿ 1 λ where + d p posz (inc q) r → absurd (r (ext refl)) + d p (possuc x) (inc q) r → inc (pos x) + +*ℚ-positive : ∀ {x y} → Positive x → Positive y → Positive (x *ℚ y) +*ℚ-positive = work _ _ where + work : ∀ x y → Positive x → Positive y → Positive (x *ℚ y) + work = ℚ-elim-propⁿ 2 λ d p a b (inc x) (inc y) → inc (ℤ.*ℤ-positive x y) + ++ℚ-positive : ∀ {x y} → Positive x → Positive y → Positive (x +ℚ y) ++ℚ-positive = work _ _ where + work : ∀ x y → Positive x → Positive y → Positive (x +ℚ y) + work = ℚ-elim-propⁿ 2 λ d p a b (inc x) (inc y) → inc (ℤ.+ℤ-positive (ℤ.*ℤ-positive x p) (ℤ.*ℤ-positive y p)) +``` diff --git a/src/Data/Set/Coequaliser.lagda.md b/src/Data/Set/Coequaliser.lagda.md index c50a763da..bc82ad9d7 100644 --- a/src/Data/Set/Coequaliser.lagda.md +++ b/src/Data/Set/Coequaliser.lagda.md @@ -348,6 +348,16 @@ inc-is-surjective (squash x y p q i j) = is-prop→squarep (λ j → inc-is-surjective (p j)) (λ j → inc-is-surjective (q j)) (λ i → inc-is-surjective y) i j + +Quot-op₂ : ∀ {C : Type ℓ} {T : C → C → Type ℓ'} + → (∀ x → R x x) → (∀ y → S y y) + → (_⋆_ : A → B → C) + → ((a b : A) (x y : B) → R a b → S x y → T (a ⋆ x) (b ⋆ y)) + → A / R → B / S → C / T +Quot-op₂ Rr Sr op resp = + Coeq-rec₂ squash (λ x y → inc (op x y)) + (λ { z (x , y , r) → quot (resp x y z z r (Sr z)) }) + λ { z (x , y , r) → quot (resp z z x y (Rr z) r) } ``` --> @@ -427,21 +437,27 @@ proof that equivalence relations are `effective`{.Agda}. ```agda effective : ∀ {x y : A} → Path quotient (inc x) (inc y) → x ∼ y effective = equiv→inverse is-effective + + reflᶜ' : ∀ {x y : A} → x ≡ y → x ∼ y + reflᶜ' {x = x} p = transport (λ i → x ∼ p i) reflᶜ + + op₂ + : (f : A → A → A) + → (∀ x y u v → x ∼ u → y ∼ v → f x y ∼ f u v) + → quotient → quotient → quotient + op₂ f r = Quot-op₂ (λ x → reflᶜ) (λ x → reflᶜ) f (λ a b x y → r a x b y) + + op₂-comm + : (f : A → A → A) + → (∀ a b → f a b ∼ f b a) + → (∀ x u v → u ∼ v → f x u ∼ f x v) + → quotient → quotient → quotient + op₂-comm f c r = op₂ f (λ x y u v p q → r x y v q ∙ᶜ c x v ∙ᶜ r v x u p ∙ᶜ c v u) ``` --> + +```agda +module Data.Set.Coequaliser.Split where +``` + +# Split quotients {defines="split-quotient"} + +Recall that a [[quotient]] of a set $A$ by an equivalence relation $x +\sim y$ allows us to "replace" the identity type of $A$ by the relation +$R$, by means of a [[surjection]] $[-] : A \epi A/\sim$ having $[x] = +[y]$ equivalent to $x \sim y$. However, depending on the specifics of +$R$, we may not need to actually take a quotient after all! It may be +the case that $A/\sim$ is equivalent to a particular *subset* of $A$. +When this is the case, we say that the quotient is **split**. This +module outlines sufficient conditions for a quotient to split, by +appealing to our intuition about *normal forms*. + + + +```agda +record is-split-congruence (R : Congruence A ℓ') : Type (level-of A ⊔ ℓ') where + open Congruence R + + field + has-is-set : is-set A +``` + +A normalisation function $n : A \to A$, for an equivalence relation $x +\sim y$, is one that is the identity "up to $\sim$", i.e., we have $x +\sim n(x)$, and which respects the relation, in that, if we have $x \sim +y$, then $n(x) = n(y)$. + +```agda + normalise : A → A + + represents : ∀ x → x ∼ normalise x + respects : ∀ x y → x ∼ y → normalise x ≡ normalise y +``` + + + +It turns out that this is just enough to asking for a splitting of the +quotient map $[-]$: We can define a function sending each $x : A/\sim$ +to a fibre of the quotient map over it, by induction. At the points of +$A$, we take the fibre over $[x]$ to be $n(x)$, and we have $[n(x)] = +[x]$ by the first assumption. By the second assumption, this procedure +respects the quotient. + +```agda + splitting : (x : Congruence.quotient R) → fibre inc x + splitting (inc x) = record { fst = normalise x ; snd = quot (symᶜ (represents x)) } + splitting (glue (x , y , r) i) = record + { fst = respects x y r i + ; snd = is-prop→pathp (λ i → Coeq.squash (inc (respects x y r i)) (quot r i)) + (quot (symᶜ (represents x))) (quot (symᶜ (represents y))) i + } + splitting (squash x y p q i j) = is-set→squarep + (λ i j → hlevel {T = fibre inc (squash x y p q i j)} 2) + (λ i → splitting x) (λ i → splitting (p i)) (λ i → splitting (q i)) (λ i → splitting y) i j +``` + + + +Two other consequences of these assumptions are that the normalisation +function is literally idempotent, and that it additionally *reflects* +the quotient relation, in that $x \sim y$ is *equivalent to* $n(x) = +n(y)$. + +```agda + normalises : ∀ x → normalise (normalise x) ≡ normalise x + normalises x = respects (normalise x) x (symᶜ (represents x)) + + reflects : ∀ x y → normalise x ≡ normalise y → x ∼ y + reflects x y p = represents x ∙ᶜ reflᶜ' p ∙ᶜ symᶜ (represents y) +``` + +Finally, we show that any splitting of the quotient map generates a +normalisation procedure in the above sense: if we have a map $c : A/\sim +\to A$, we define the normalisation procedure to be $n(x) = c[x]$. + +```agda +split-surjection→is-split-congruence + : ⦃ _ : H-Level A 2 ⦄ (R : Congruence A ℓ) + → (∀ x → fibre {B = Congruence.quotient R} inc x) + → is-split-congruence R +split-surjection→is-split-congruence R split = record + { has-is-set = hlevel 2 + ; normalise = λ x → split (inc x) .fst + ; represents = λ x → effective (sym (split (inc x) .snd)) + ; respects = λ x y p → ap (fst ∘ split) (quot p) + } where open Congruence R +``` diff --git a/src/Data/Set/Material.lagda.md b/src/Data/Set/Material.lagda.md index c4544b8ca..a02f35809 100644 --- a/src/Data/Set/Material.lagda.md +++ b/src/Data/Set/Material.lagda.md @@ -520,7 +520,7 @@ constructor to the successor set. ```agda ℕV : V ℓ - ℕV = set (Lift ℓ Nat) λ x → go (Lift.lower x) where + ℕV = set (Lift ℓ Nat) λ x → go (lower x) where go : Nat → V ℓ go zero = ∅V go (suc x) = suc-V (go x) @@ -558,7 +558,7 @@ set. (inr w) → do (pred , ix , w) ← w (ix , x) ← ix - pure (lift (suc (ix .Lift.lower)) , ap₂ _∪V_ x (ap singleton x) ∙ sym w)) + pure (lift (suc (ix .lower)) , ap₂ _∪V_ x (ap singleton x) ∙ sym w)) ``` ## Replacement diff --git a/src/Data/Sum/Properties.lagda.md b/src/Data/Sum/Properties.lagda.md index bab379dec..7a50bccd2 100644 --- a/src/Data/Sum/Properties.lagda.md +++ b/src/Data/Sum/Properties.lagda.md @@ -73,8 +73,8 @@ the `Code`{.Agda} computes to `the empty type`{.Agda ident=⊥}. ```agda decode : {x y : A ⊎ B} → Code x y → x ≡ y - decode {x = inl x} {y = inl x₁} code = ap inl (Lift.lower code) - decode {x = inr x} {y = inr x₁} code = ap inr (Lift.lower code) + decode {x = inl x} {y = inl x₁} code = ap inl (lower code) + decode {x = inr x} {y = inr x₁} code = ap inr (lower code) ``` In the inverse direction, we have a procedure for turning paths into diff --git a/src/Order/DCPO/Free.lagda.md b/src/Order/DCPO/Free.lagda.md index 89334d4bd..9f6558174 100644 --- a/src/Order/DCPO/Free.lagda.md +++ b/src/Order/DCPO/Free.lagda.md @@ -281,7 +281,7 @@ $$. ```agda part-counit : ↯ Ob → Ob - part-counit x = ⋃-prop (x .elt ⊙ Lift.lower) def-prop where abstract + part-counit x = ⋃-prop (x .elt ⊙ lower) def-prop where abstract def-prop : is-prop (Lift o ⌞ x ⌟) def-prop = hlevel 1 ``` @@ -298,7 +298,7 @@ And if $x$ is undefined, then this is the bottom element. part-counit-¬elt : (x : ↯ Ob) → (⌞ x ⌟ → ⊥) → part-counit x ≡ bottom part-counit-¬elt x ¬def = ≤-antisym - (⋃-prop-least _ _ _ (λ p → absurd (¬def (Lift.lower p)))) + (⋃-prop-least _ _ _ (λ p → absurd (¬def (lower p)))) (bottom≤x _) ``` @@ -324,7 +324,7 @@ curious reader. part-counit-⊑ {x = x} {y = y} p = ⋃-prop-least _ _ (part-counit y) λ (lift i) → x .elt i =˘⟨ p .refines i ⟩ y .elt (p .implies i) ≤⟨ ⋃-prop-le _ _ (lift (p .implies i)) ⟩ - ⋃-prop (y .elt ⊙ Lift.lower) _ ≤∎ + ⋃-prop (y .elt ⊙ lower) _ ≤∎ part-counit-lub s sdir .is-lub.fam≤lub i = ⋃-prop-least _ _ _ λ (lift p) → @@ -358,10 +358,10 @@ Free-Pointed-dcpo⊣Forget-Pointed-dcpo .counit .is-natural D E f = ext λ x → sym $ Strict-scott.pres-⋃-prop f _ _ _ Free-Pointed-dcpo⊣Forget-Pointed-dcpo .zig {A} = ext λ x → part-ext - (A?.⋃-prop-least _ _ x (λ p → always-⊒ (Lift.lower p , refl)) .implies) + (A?.⋃-prop-least _ _ x (λ p → always-⊒ (lower p , refl)) .implies) (λ p → A?.⋃-prop-le _ _ (lift p) .implies tt) (λ p q → - sym (A?.⋃-prop-least _ _ x (λ p → always-⊒ (Lift.lower p , refl)) .refines p) + sym (A?.⋃-prop-least _ _ x (λ p → always-⊒ (lower p , refl)) .refines p) ∙ ↯-indep x) where module A? = Pointed-dcpo (Parts-pointed-dcpo A) diff --git a/src/Order/DCPO/Pointed.lagda.md b/src/Order/DCPO/Pointed.lagda.md index a4f4df9a5..aa373fee2 100644 --- a/src/Order/DCPO/Pointed.lagda.md +++ b/src/Order/DCPO/Pointed.lagda.md @@ -530,7 +530,7 @@ Scott-continuous. pres-bot : ∀ x → is-bottom D.poset x → is-bottom E.poset (f x) pres-bot x x-bot y = f x E.≤⟨ monotone (x-bot _) ⟩ - f (D.⋃-semi _ _) E.≤⟨ is-lub.least (pres-⋃-semi (λ x → absurd (x .Lift.lower)) (λ ())) y (λ ()) ⟩ + f (D.⋃-semi _ _) E.≤⟨ is-lub.least (pres-⋃-semi (λ x → absurd (x .lower)) (λ ())) y (λ ()) ⟩ y E.≤∎ ``` @@ -567,6 +567,6 @@ families, then $f$ must be monotonic, and thus strictly Scott-continuous. (to-scott-directed f (λ s dir x lub → pres s (is-directed-family.semidirected dir) x lub)) (λ x x-bot y → is-lub.least - (pres _ (λ x → absurd (x .Lift.lower)) x (lift-is-lub (is-bottom→is-lub D.poset {f = λ ()} x-bot))) + (pres _ (λ x → absurd (x .lower)) x (lift-is-lub (is-bottom→is-lub D.poset {f = λ ()} x-bot))) y (λ ())) ``` diff --git a/src/Order/Diagram/Glb.lagda.md b/src/Order/Diagram/Glb.lagda.md index 3d1deddd5..38c58654a 100644 --- a/src/Order/Diagram/Glb.lagda.md +++ b/src/Order/Diagram/Glb.lagda.md @@ -92,25 +92,25 @@ module _ {o ℓ} {P : Poset o ℓ} where lift-is-glb : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} {glb} - → is-glb P F glb → is-glb P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) glb + → is-glb P F glb → is-glb P (F ⊙ lower {ℓ = ℓᵢ'}) glb lift-is-glb is .glb≤fam (lift ix) = is .glb≤fam ix lift-is-glb is .greatest ub' le = is .greatest ub' (le ⊙ lift) lift-glb : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} - → Glb P F → Glb P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) + → Glb P F → Glb P (F ⊙ lower {ℓ = ℓᵢ'}) lift-glb glb .Glb.glb = Glb.glb glb lift-glb glb .Glb.has-glb = lift-is-glb (Glb.has-glb glb) lower-is-glb : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} {glb} - → is-glb P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) glb → is-glb P F glb + → is-glb P (F ⊙ lower {ℓ = ℓᵢ'}) glb → is-glb P F glb lower-is-glb is .glb≤fam ix = is .glb≤fam (lift ix) - lower-is-glb is .greatest ub' le = is .greatest ub' (le ⊙ Lift.lower) + lower-is-glb is .greatest ub' le = is .greatest ub' (le ⊙ lower) lower-glb : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} - → Glb P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) → Glb P F + → Glb P (F ⊙ lower {ℓ = ℓᵢ'}) → Glb P F lower-glb glb .Glb.glb = Glb.glb glb lower-glb glb .Glb.has-glb = lower-is-glb (Glb.has-glb glb) ``` diff --git a/src/Order/Diagram/Lub.lagda.md b/src/Order/Diagram/Lub.lagda.md index 96f1c7967..78f03be20 100644 --- a/src/Order/Diagram/Lub.lagda.md +++ b/src/Order/Diagram/Lub.lagda.md @@ -90,25 +90,25 @@ module _ {o ℓ} {P : Poset o ℓ} where lift-is-lub : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} {lub} - → is-lub P F lub → is-lub P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) lub + → is-lub P F lub → is-lub P (F ⊙ lower {ℓ = ℓᵢ'}) lub lift-is-lub is .fam≤lub (lift ix) = is .fam≤lub ix lift-is-lub is .least ub' le = is .least ub' (le ⊙ lift) lift-lub : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} - → Lub P F → Lub P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) + → Lub P F → Lub P (F ⊙ lower {ℓ = ℓᵢ'}) lift-lub lub .Lub.lub = Lub.lub lub lift-lub lub .Lub.has-lub = lift-is-lub (Lub.has-lub lub) lower-is-lub : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} {lub} - → is-lub P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) lub → is-lub P F lub + → is-lub P (F ⊙ lower {ℓ = ℓᵢ'}) lub → is-lub P F lub lower-is-lub is .fam≤lub ix = is .fam≤lub (lift ix) - lower-is-lub is .least ub' le = is .least ub' (le ⊙ Lift.lower) + lower-is-lub is .least ub' le = is .least ub' (le ⊙ lower) lower-lub : ∀ {ℓᵢ ℓᵢ'} {I : Type ℓᵢ} {F : I → Ob} - → Lub P (F ⊙ Lift.lower {ℓ = ℓᵢ'}) → Lub P F + → Lub P (F ⊙ lower {ℓ = ℓᵢ'}) → Lub P F lower-lub lub .Lub.lub = Lub.lub lub lower-lub lub .Lub.has-lub = lower-is-lub (Lub.has-lub lub) ``` diff --git a/src/Order/Instances/Lift.lagda.md b/src/Order/Instances/Lift.lagda.md index b60cbc6d9..caf76276f 100644 --- a/src/Order/Instances/Lift.lagda.md +++ b/src/Order/Instances/Lift.lagda.md @@ -20,7 +20,7 @@ Fortunately you can uniformly lift the poset to this bigger universe. ```agda Liftᵖ : ∀ {o r} o′ r′ → Poset o r → Poset (o ⊔ o′) (r ⊔ r′) Liftᵖ o' r' X .Poset.Ob = Lift o' ⌞ X ⌟ -Liftᵖ o' r' X .Poset._≤_ x y = Lift r' $ (X .Poset._≤_) (Lift.lower x) (Lift.lower y) +Liftᵖ o' r' X .Poset._≤_ x y = Lift r' $ (X .Poset._≤_) (lower x) (lower y) Liftᵖ o' r' X .Poset.≤-thin = Lift-is-hlevel 1 $ X .Poset.≤-thin Liftᵖ o' r' X .Poset.≤-refl = lift $ X .Poset.≤-refl Liftᵖ o' r' X .Poset.≤-trans (lift p) (lift q) = lift (X .Poset.≤-trans p q) diff --git a/src/Order/Semilattice/Join/Reasoning.lagda.md b/src/Order/Semilattice/Join/Reasoning.lagda.md index a6ae3a657..a2b1210de 100644 --- a/src/Order/Semilattice/Join/Reasoning.lagda.md +++ b/src/Order/Semilattice/Join/Reasoning.lagda.md @@ -11,6 +11,8 @@ open import Order.Base import Cat.Reasoning import Order.Reasoning + +open is-monoid ``` --> diff --git a/src/Order/Semilattice/Meet/Reasoning.lagda.md b/src/Order/Semilattice/Meet/Reasoning.lagda.md index 438e79baf..e8daf4581 100644 --- a/src/Order/Semilattice/Meet/Reasoning.lagda.md +++ b/src/Order/Semilattice/Meet/Reasoning.lagda.md @@ -11,6 +11,8 @@ open import Order.Base import Cat.Reasoning import Order.Reasoning + +open is-monoid ``` --> diff --git a/src/Topoi/Base.lagda.md b/src/Topoi/Base.lagda.md index cb971faed..9cc78ae50 100644 --- a/src/Topoi/Base.lagda.md +++ b/src/Topoi/Base.lagda.md @@ -1,26 +1,34 @@ We start by defining a "global" product-assigning operation. ```agda - _⊗₀_ : Ob → Ob → Ob - a ⊗₀ b = product.apex {a} {b} + module _ a b where + open Product (all-products a b) + renaming (apex to infixr 7 _⊗₀_) + using () public ``` This operation extends to a bifunctor $\cC \times \cC \to \cC$. diff --git a/src/Order/DCPO.lagda.md b/src/Order/DCPO.lagda.md index 553e8ee45..0ef196b6a 100644 --- a/src/Order/DCPO.lagda.md +++ b/src/Order/DCPO.lagda.md @@ -237,7 +237,7 @@ module DCPOs {o ℓ : Level} = Cat.Reasoning (DCPOs o ℓ) DCPO : (o ℓ : Level) → Type _ DCPO o ℓ = DCPOs.Ob {o} {ℓ} -DCPOs↪Posets : ∀ {o ℓ} → Functor (DCPOs o ℓ) (Posets o ℓ) +DCPOs↪Posets : ∀ {o ℓ} → Functor (DCPOs o ℓ) (Posets o ℓ) DCPOs↪Posets = Forget-subcat DCPOs↪Sets : ∀ {o ℓ} → Functor (DCPOs o ℓ) (Sets o) @@ -261,7 +261,7 @@ module DCPO {o ℓ} (D : DCPO o ℓ) where poset : Poset o ℓ poset = D .fst - open Order.Reasoning poset public + open Order.Reasoning (D .fst) public set : Set o set = el ⌞ D ⌟ Ob-is-set @@ -269,7 +269,7 @@ module DCPO {o ℓ} (D : DCPO o ℓ) where has-dcpo : is-dcpo poset has-dcpo = D .snd - open is-dcpo has-dcpo public + open is-dcpo (D .snd) public ⋃-pointwise : ∀ {Ix} {s s' : Ix → Ob} From d195da5c5f8cb0363cab162aeaea865dd8d26407 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Am=C3=A9lia?= Date: Wed, 18 Sep 2024 13:15:32 -0300 Subject: [PATCH 9/9] =?UTF-8?q?Solver=20for=20equations=20over=20=E2=84=9A?= =?UTF-8?q?=20(#431)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We can re-use the ring solver for expressions over ℚ, even involving inverses. The gist of it is that we can re-express any equation involving inverses to an equation between quotients of polynomials, then invoke the ring solver. Even though this idea seemed pretty dumb (surely there would be a cleverer algorithm over ℚ?), it's what the Rocq `field` tactic does as well. I also added some more comments explicating why ℚ is defined the way it is. --- src/1Lab/HLevel/Closure.lagda.md | 14 + src/Algebra/Group/Instances/Integers.lagda.md | 2 +- src/Algebra/Ring/Solver.agda | 13 +- src/Data/Fin/Product.lagda.md | 17 + src/Data/Int/DivMod.lagda.md | 2 +- src/Data/Int/Order.lagda.md | 203 ++++++++++++ src/Data/Int/Properties.lagda.md | 8 + src/Data/Nat/Order.lagda.md | 9 + src/Data/Nat/Properties.lagda.md | 8 + src/Data/Rational/Base.lagda.md | 106 +++++- src/Data/Rational/Order.lagda.md | 253 +++++++++++--- src/Data/Rational/Properties.lagda.md | 224 +++++++++++++ src/Data/Rational/Reflection.agda | 55 ++++ src/Data/Rational/Solver.agda | 310 ++++++++++++++++++ src/Prim/Data/Sigma.lagda.md | 4 + 15 files changed, 1170 insertions(+), 58 deletions(-) create mode 100644 src/Data/Rational/Properties.lagda.md create mode 100644 src/Data/Rational/Reflection.agda create mode 100644 src/Data/Rational/Solver.agda diff --git a/src/1Lab/HLevel/Closure.lagda.md b/src/1Lab/HLevel/Closure.lagda.md index d626fec0f..19e613658 100644 --- a/src/1Lab/HLevel/Closure.lagda.md +++ b/src/1Lab/HLevel/Closure.lagda.md @@ -188,6 +188,14 @@ between functions is a function into identities: (λ f {x} → f x) (λ f x → f) (λ _ → refl) (Π-is-hlevel n bhl) +Π-is-hlevel-inst + : ∀ {a b} {A : Type a} {B : A → Type b} + → (n : Nat) (Bhl : (x : A) → is-hlevel (B x) n) + → is-hlevel (⦃ x : A ⦄ → B x) n +Π-is-hlevel-inst n bhl = retract→is-hlevel n + (λ f ⦃ x ⦄ → f x) (λ f x → f ⦃ x ⦄) (λ _ → refl) + (Π-is-hlevel n bhl) + Π-is-hlevel² : ∀ {a b c} {A : Type a} {B : A → Type b} {C : ∀ a → B a → Type c} → (n : Nat) (Bhl : (x : A) (y : B x) → is-hlevel (C x y) n) @@ -376,6 +384,12 @@ instance opaque → H-Level (∀ {x} → S x) n H-Level-pi' {n = n} .H-Level.has-hlevel = Π-is-hlevel' n λ _ → hlevel n + H-Level-pi'' + : ∀ {n} {S : T → Type ℓ} + → ⦃ ∀ {x} → H-Level (S x) n ⦄ + → H-Level (∀ ⦃ x ⦄ → S x) n + H-Level-pi'' {n = n} .H-Level.has-hlevel = Π-is-hlevel-inst n λ _ → hlevel n + H-Level-sigma : ∀ {n} {S : T → Type ℓ} → ⦃ H-Level T n ⦄ → ⦃ ∀ {x} → H-Level (S x) n ⦄ diff --git a/src/Algebra/Group/Instances/Integers.lagda.md b/src/Algebra/Group/Instances/Integers.lagda.md index 23df0c39e..677147504 100644 --- a/src/Algebra/Group/Instances/Integers.lagda.md +++ b/src/Algebra/Group/Instances/Integers.lagda.md @@ -9,7 +9,7 @@ open import Cat.Prelude open import Data.Int.Universal open import Data.Nat.Order -open import Data.Int hiding (Positive) +open import Data.Int hiding (Positive ; <-not-equal) open import Data.Nat open is-group-hom diff --git a/src/Algebra/Ring/Solver.agda b/src/Algebra/Ring/Solver.agda index 8999382b8..33cec2c3d 100644 --- a/src/Algebra/Ring/Solver.agda +++ b/src/Algebra/Ring/Solver.agda @@ -46,10 +46,15 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where ℤ↪R-rh = Int-is-initial R' .centre module ℤ↪R = is-ring-hom (ℤ↪R-rh .preserves) - embed-coe = ℤ↪R-rh .hom open CRing-on cring using (*-commutes) + embed-coe : Int → R + embed-coe x = ℤ↪R-rh .hom (lift x) + + embed-lemma : {h' : Int → R} → is-ring-hom (Liftℤ {ℓ} .snd) (R' .snd) (h' ⊙ lower) → ∀ x → embed-coe x ≡ h' x + embed-lemma p x = happly (ap Total-hom.hom (Int-is-initial R' .paths (total-hom _ p))) (lift x) + data Poly : Nat → Type ℓ data Normal : Nat → Type ℓ @@ -67,7 +72,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where Ep ∅ i = R.0r Ep (p *x+ c) (x ∷ e) = Ep p (x ∷ e) R.* x R.+ En c e - En (con x) i = embed-coe (lift x) + En (con x) i = embed-coe x En (poly x) i = Ep x i 0h : ∀ {n} → Poly n @@ -200,7 +205,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where Number-Polynomial .Number.fromNat n = con (pos n) eval (op o p₁ p₂) ρ = sem o (⟦ p₁ ⟧ ρ) (⟦ p₂ ⟧ ρ) - eval (con c) ρ = embed-coe (lift c) + eval (con c) ρ = embed-coe c eval (var x) ρ = lookup ρ x eval (:- p) ρ = R.- ⟦ p ⟧ ρ @@ -371,7 +376,7 @@ module Impl {ℓ} {R : Type ℓ} (cring : CRing-on R) where -ₙ-hom (poly x) ρ = -ₚ-hom x ρ sound-coe - : ∀ {n} (c : Int) (ρ : Vec R n) → En (normal-coe c) ρ ≡ embed-coe (lift c) + : ∀ {n} (c : Int) (ρ : Vec R n) → En (normal-coe c) ρ ≡ embed-coe c sound-coe c [] = refl sound-coe c (x ∷ ρ) = ∅*x+ₙ-hom (normal-coe c) x ρ ∙ sound-coe c ρ diff --git a/src/Data/Fin/Product.lagda.md b/src/Data/Fin/Product.lagda.md index a0d1403c7..ee06b37e8 100644 --- a/src/Data/Fin/Product.lagda.md +++ b/src/Data/Fin/Product.lagda.md @@ -193,6 +193,23 @@ curry-∀ᶠ → ∀ᶠ n P Q curry-∀ᶠ {zero} f = f tt curry-∀ᶠ {suc n} f a = curry-∀ᶠ {n} λ b → f (a , b) + +∀ᶠⁱ : ∀ n {ℓ ℓ'} (P : (i : Fin n) → Type (ℓ i)) (Q : Πᶠ P → Type ℓ') → Type (ℓ-maxᶠ ℓ ⊔ ℓ') +∀ᶠⁱ zero P Q = Q tt +∀ᶠⁱ (suc n) P Q = {a : P fzero} → ∀ᶠⁱ n (λ i → P (fsuc i)) (λ b → Q (a , b)) + +apply-∀ᶠⁱ + : ∀ {n} {ℓ : Fin n → Level} {ℓ'} {P : (i : Fin n) → Type (ℓ i)} {Q : Πᶠ P → Type ℓ'} + → ∀ᶠⁱ n P Q → (a : Πᶠ P) → Q a +apply-∀ᶠⁱ {zero} f a = f +apply-∀ᶠⁱ {suc n} f (a , as) = apply-∀ᶠⁱ (f {a}) as + +curry-∀ᶠⁱ + : ∀ {n} {ℓ : Fin n → Level} {ℓ'} {P : (i : Fin n) → Type (ℓ i)} {Q : Πᶠ P → Type ℓ'} + → ((a : Πᶠ P) → Q a) + → ∀ᶠⁱ n P Q +curry-∀ᶠⁱ {zero} f = f tt +curry-∀ᶠⁱ {suc n} f {a} = curry-∀ᶠⁱ {n} λ b → f (a , b) ``` --> diff --git a/src/Data/Int/DivMod.lagda.md b/src/Data/Int/DivMod.lagda.md index ab758f8ad..06c81abcf 100644 --- a/src/Data/Int/DivMod.lagda.md +++ b/src/Data/Int/DivMod.lagda.md @@ -10,7 +10,7 @@ open import Data.Int.Divisible open import Data.Nat.DivMod open import Data.Dec.Base open import Data.Fin hiding (_<_) -open import Data.Int hiding (Positive) +open import Data.Int hiding (Positive ; _<_ ; <-weaken) open import Data.Nat as Nat ``` --> diff --git a/src/Data/Int/Order.lagda.md b/src/Data/Int/Order.lagda.md index 2a101b4d2..97fa46bf0 100644 --- a/src/Data/Int/Order.lagda.md +++ b/src/Data/Int/Order.lagda.md @@ -7,6 +7,7 @@ open import 1Lab.Type open import Data.Int.Properties open import Data.Int.Base +open import Data.Sum.Base open import Data.Dec import Data.Nat.Properties as Nat @@ -33,6 +34,20 @@ data _≤_ : Int → Int → Type where neg≤pos : ∀ {x y} → negsuc x ≤ pos y ``` + + Note the key properties: the ordering between negative numbers is reversed, and every negative number is smaller than every positive number. This means that $\bZ$ decomposes, as an order, into an _ordinal @@ -245,4 +260,192 @@ abstract *ℤ-nonnegative : ∀ {x y} → 0 ≤ x → 0 ≤ y → 0 ≤ (x *ℤ y) *ℤ-nonnegative {pos x} {pos y} (pos≤pos p) (pos≤pos q) = ≤-trans (pos≤pos Nat.0≤x) (≤-refl' (sym (assign-pos (x * y)))) + + positive→nonnegative : ∀ {x} → Positive x → 0 ≤ x + positive→nonnegative (pos x) = pos≤pos Nat.0≤x + + -ℕ-nonnegative : ∀ {x y} → y Nat.≤ x → 0 ≤ (x ℕ- y) + -ℕ-nonnegative {x} {y} Nat.0≤x = pos≤pos Nat.0≤x + -ℕ-nonnegative {suc x} {suc y} (Nat.s≤s p) = -ℕ-nonnegative p + + -ℤ-nonnegative : ∀ {x y} → 0 ≤ x → 0 ≤ y → y ≤ x → 0 ≤ (x -ℤ y) + -ℤ-nonnegative {posz} {posz} (pos≤pos p) (pos≤pos q) (pos≤pos r) = pos≤pos Nat.0≤x + -ℤ-nonnegative {possuc x} {posz} (pos≤pos p) (pos≤pos q) (pos≤pos r) = pos≤pos Nat.0≤x + -ℤ-nonnegative {possuc x} {possuc y} (pos≤pos p) (pos≤pos q) (pos≤pos r) = -ℕ-nonnegative (Nat.≤-peel r) +``` + +# The strict order + +```agda +data _<_ : Int → Int → Type where + pos @@ -562,7 +604,7 @@ common-denominator (suc sz) fs with (c , c≠0 , nums , prfs) ← common-denomin ℚ-elim-propⁿ : ∀ (n : Nat) {ℓ} {P : Arrᶠ {n = n} (λ _ → ℚ) (Type ℓ)} → ⦃ _ : {as : Πᶠ (λ i → ℚ)} → H-Level (applyᶠ P as) 1 ⦄ - → ( (d : Int) (p : Positive d) + → ( (d : Int) ⦃ p : Positive d ⦄ → ∀ᶠ n (λ i → Int) (λ as → applyᶠ P (mapₚ (λ i n → toℚ (n / d [ p ])) as))) → ∀ᶠ n (λ _ → ℚ) λ as → applyᶠ P as @@ -588,7 +630,7 @@ common-denominator (suc sz) fs with (c , c≠0 , nums , prfs) ← common-denomin rats = mapₚ (λ i n → toℚ (n / d [ d≠0 ])) (tabulateₚ nums) p₀ : applyᶠ P rats - p₀ = apply-∀ᶠ (work d d≠0) (tabulateₚ nums) + p₀ = apply-∀ᶠ (work d ⦃ d≠0 ⦄) (tabulateₚ nums) rats=as : rats ≡ as rats=as = extₚ λ i → @@ -620,7 +662,63 @@ instance Extensional-ℚ .reflᵉ = ℚ-elim-prop (λ _ → hlevel 1) λ { record{} → refl } Extensional-ℚ .idsᵉ .to-path {a} {b} = go a b where go : ∀ a b → ⌞ eqℚ a b ⌟ → a ≡ b - go = ℚ-elim-propⁿ 2 (λ d _ a b p → quotℚ (to-same-rational p)) + go = ℚ-elim-propⁿ 2 (λ d a b p → quotℚ (to-same-rational p)) Extensional-ℚ .idsᵉ .to-path-over p = prop! + +record Nonzero (x : ℚ) : Type where + constructor inc + -- It's important that Nonzero is a strict proposition, living in + -- Type, so that it doesn't matter what instance gets selected when we + -- use it in invℚ etc. + -- + -- The alternative is to always use it as an irrelevant instance (or + -- to, god forbid, have it in Agda's Prop), but that doesn't play well + -- with the overlap pragmas; and we need those if we're gonna have + -- e.g. Nonzero (p * q) as an instance. + field + .lower : x ≠ 0 + +instance + H-Level-Nonzero : ∀ {x n} → H-Level (Nonzero x) (suc n) + H-Level-Nonzero = prop-instance λ _ _ → refl + +open Nonzero + +abstract + from-nonzero : ∀ {x} ⦃ p : Nonzero x ⦄ → x ≠ 0 + from-nonzero ⦃ inc α ⦄ p = absurd (α p) + + from-nonzero-frac : ∀ {x y} {p : Positive y} → Nonzero (toℚ (x / y [ p ])) → x ≠ 0 + from-nonzero-frac (inc α) p = absurd (α (quotℚ (to-same-rational (*ℤ-oner _ ∙ p)))) + + to-nonzero-frac : ∀ {x y} {p : Positive y} → x ≠ 0 → Nonzero (toℚ (x / y [ p ])) + to-nonzero-frac p = inc λ α → p (sym (*ℤ-oner _) ∙ from-same-rational (unquotℚ α)) + +instance + Nonzero-neg : ∀ {x d} {p : Positive d} → Nonzero (toℚ (negsuc x / d [ p ])) + Nonzero-neg = inc (λ p → absurd (negsuc≠pos (from-same-rational (unquotℚ p)))) + + Nonzero-pos : ∀ {x d} {p : Positive d} ⦃ _ : Positive x ⦄ → Nonzero (toℚ (x / d [ p ])) + Nonzero-pos {.(possuc x)} ⦃ pos x ⦄ = inc (λ p → absurd (suc≠zero (pos-injective (from-same-rational (unquotℚ p))))) + {-# OVERLAPPABLE Nonzero-pos #-} + +-- Since we want invℚ to be injective as well, we re-wrap the result on +-- the RHS, to make sure the clause is headed by a constructor. +invℚ : (x : ℚ) ⦃ p : Nonzero x ⦄ → ℚ +invℚ (inc x) ⦃ inc α ⦄ = inc (unℚ (inverseℚ (inc x) (λ p → absurd (α p)) .fst)) + +*ℚ-invr : {x : ℚ} {p : Nonzero x} → x *ℚ invℚ x ⦃ p ⦄ ≡ 1 +*ℚ-invr {inc x} {inc α} with inverseℚ (inc x) (λ p → absurd (α p)) +... | (inc x , p) = p + +-ℚ-def : ∀ x y → x +ℚ (-ℚ y) ≡ x -ℚ y +-ℚ-def (inc x) (inc y) = refl + +_/ℚ_ : (x y : ℚ) ⦃ p : Nonzero y ⦄ → ℚ +inc x /ℚ inc y = inc x *ℚ invℚ (inc y) + +abstract + from-same-denom : ∀ {x y d} ⦃ p : Positive d ⦄ → x / d ≡ y / d → x ≡ y + from-same-denom {x} {y} {d} p = *ℤ-injectiver d x y (positive→nonzero auto) (from-same-rational (unquotℚ p)) ``` --> diff --git a/src/Data/Rational/Order.lagda.md b/src/Data/Rational/Order.lagda.md index 120af8de6..18fa0a881 100644 --- a/src/Data/Rational/Order.lagda.md +++ b/src/Data/Rational/Order.lagda.md @@ -5,9 +5,12 @@ open import 1Lab.Prelude open import Algebra.Ring.Commutative open import Algebra.Ring.Solver +open import Data.Rational.Reflection open import Data.Set.Coequaliser hiding (_/_) +open import Data.Sum.Properties open import Data.Rational.Base open import Data.Int.Base hiding (Positive ; H-Level-Positive ; Dec-Positive) +open import Data.Sum.Base open import Data.Dec open import Order.Instances.Int @@ -26,6 +29,8 @@ module Data.Rational.Order where @@ -121,36 +126,37 @@ can prove that the ordering on the rational numbers is decidable, reflexive, transitive, and antisymmetric. ```agda -≤-dec : ∀ x y → Dec (x ≤ y) -≤-dec = elim! go where - go : ∀ x y → Dec (toℚ x ≤ toℚ y) - go x@record{} y@record{} with holds? ((↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x)) - ... | yes p = yes (inc p) - ... | no ¬p = no (¬p ∘ lower) instance Dec-≤ : ∀ {x y} → Dec (x ≤ y) - Dec-≤ {x} {y} = ≤-dec x y + Dec-≤ {inc x} {inc y} = elim! {P = λ x → ∀ y → Dec (ℚ.inc x ≤ inc y)} go x y where + go : ∀ x y → Dec (toℚ x ≤ toℚ y) + go x@record{} y@record{} with holds? ((↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x)) + ... | yes p = yes (inc p) + ... | no ¬p = no (¬p ∘ lower) + + ≤-from : ∀ {x y} ⦃ p : (↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x) ⦄ → toℚ x ≤ toℚ y + ≤-from {record{}} {record{}} ⦃ p ⦄ = inc p abstract toℚ≤ : ∀ {x y} → (↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x) → toℚ x ≤ toℚ y toℚ≤ {record{}} {record{}} p = inc p - ≤-refl : ∀ {x} → x ≤ x - ≤-refl {x} = work x where - work : ∀ x → x ≤ x - work = elim! λ f → toℚ≤ ℤ.≤-refl + ≤-refl : ∀ {x} → x ≤ x + ≤-trans : ∀ {x y z} → x ≤ y → y ≤ z → x ≤ z + ≤-antisym : ∀ {x y} → x ≤ y → y ≤ x → x ≡ y + + unquoteDef ≤-refl ≤-trans ≤-antisym = do + by-elim-ℚ ≤-refl λ d x → inc ℤ.≤-refl - ≤-trans : ∀ {x y z} → x ≤ y → y ≤ z → x ≤ z - ≤-trans {x} {y} {z} = work x y z where - work : ∀ x y z → x ≤ y → y ≤ z → x ≤ z - work = elim! λ x y z p q → inc (≤ᶠ-trans x y z (p .lower) (q .lower)) + by-elim-ℚ ≤-trans λ d x y z (inc p) (inc q) → inc + (≤ᶠ-trans (x / d [ auto ]) (y / d [ auto ]) (z / d [ auto ]) p q) - ≤-antisym : ∀ {x y} → x ≤ y → y ≤ x → x ≡ y - ≤-antisym {x} {y} = work x y where - work : ∀ x y → x ≤ y → y ≤ x → x ≡ y - work = elim! λ where - x@record{} y@record{} p q → quotℚ (to-same-rational (ℤ.≤-antisym (p .lower) (q .lower))) + by-elim-ℚ ≤-antisym λ d x y (inc α) (inc β) → + quotℚ (to-same-rational (ℤ.≤-antisym α β)) + + ≤-refl' : ∀ {x y} (p : x ≡ y) → x ≤ y + ≤-refl' {x} p = transport (λ i → x ≤ p i) ≤-refl ``` ## Algebraic properties @@ -164,19 +170,31 @@ order. common-denominator-≤ : ∀ {x y d} (p : ℤ.Positive d) → x ℤ.≤ y → toℚ (x / d [ p ]) ≤ toℚ (y / d [ p ]) common-denominator-≤ {d = d} (pos _) p = inc (ℤ.*ℤ-preserves-≤r d p) - +ℚ-preserves-≤ : ∀ x y a b → x ≤ y → a ≤ b → (x +ℚ a) ≤ (y +ℚ b) - +ℚ-preserves-≤ = ℚ-elim-propⁿ 4 λ d d≠0 x y a b (inc p) (inc q) → - common-denominator-≤ (ℤ.*ℤ-positive d≠0 d≠0) $ - x *ℤ d +ℤ a *ℤ d ≤⟨ ℤ.+ℤ-preserves-≤r (a *ℤ d) _ _ p ⟩ - y *ℤ d +ℤ a *ℤ d ≤⟨ ℤ.+ℤ-preserves-≤l (y *ℤ d) _ _ q ⟩ - y *ℤ d +ℤ b *ℤ d ≤∎ + +ℚ-preserves-≤ : ∀ {x y a b} → x ≤ y → a ≤ b → (x +ℚ a) ≤ (y +ℚ b) + *ℚ-nonnegative : ∀ {x y} → 0 ≤ x → 0 ≤ y → 0 ≤ (x *ℚ y) + invℚ-nonnegative : ∀ {x} ⦃ p : Nonzero x ⦄ → 0 ≤ x → 0 ≤ (invℚ x ⦃ p ⦄) - *ℚ-nonnegative : ∀ x y → 0 ≤ x → 0 ≤ y → 0 ≤ (x *ℚ y) - *ℚ-nonnegative = ℚ-elim-propⁿ 2 λ d d≠0 x y (inc p) (inc q) → inc (ℤ.≤-trans - (ℤ.*ℤ-nonnegative - (ℤ.≤-trans p (ℤ.≤-refl' (ℤ.*ℤ-oner x))) - (ℤ.≤-trans q (ℤ.≤-refl' (ℤ.*ℤ-oner y)))) - (ℤ.≤-refl' (sym (ℤ.*ℤ-oner (x *ℤ y))))) + unquoteDef +ℚ-preserves-≤ *ℚ-nonnegative invℚ-nonnegative = do + + by-elim-ℚ +ℚ-preserves-≤ λ d x y a b (inc p) (inc q) → + common-denominator-≤ (ℤ.*ℤ-positive auto auto) $ + x *ℤ d +ℤ a *ℤ d ≤⟨ ℤ.+ℤ-preserves-≤r (a *ℤ d) _ _ p ⟩ + y *ℤ d +ℤ a *ℤ d ≤⟨ ℤ.+ℤ-preserves-≤l (y *ℤ d) _ _ q ⟩ + y *ℤ d +ℤ b *ℤ d ≤∎ + + by-elim-ℚ *ℚ-nonnegative λ d x y (inc p) (inc q) → inc + (ℤ.≤-trans + (ℤ.*ℤ-nonnegative + (ℤ.≤-trans p (ℤ.≤-refl' (ℤ.*ℤ-oner x))) + (ℤ.≤-trans q (ℤ.≤-refl' (ℤ.*ℤ-oner y)))) + (ℤ.≤-refl' (sym (ℤ.*ℤ-oner (x *ℤ y))))) + + by-elim-ℚ invℚ-nonnegative λ where + (possuc x) posz ⦃ inc nz ⦄ z → absurd (nz (quotℚ (to-same-rational refl))) + (possuc x) (possuc y) z → inc (ℤ.pos≤pos 0≤x) + + /ℚ-nonnegative : ∀ {x y} ⦃ p : Nonzero y ⦄ → 0 ≤ x → 0 ≤ y → 0 ≤ (x /ℚ y) + /ℚ-nonnegative {inc x} {inc y} a b = *ℚ-nonnegative a (invℚ-nonnegative {inc y} b) ``` ## Positivity @@ -209,7 +227,7 @@ instance Dec-Positive : ∀ {x} → Dec (Positive x) Dec-Positive {x} with (r@(n / d [ p ]) , q) ← splitℚ x | holds? (ℤ.Positive n) ... | yes p = yes (subst Positive q (inc (recover p))) - ... | no ¬p = no λ x → case subst Positive (sym q) x of λ (inc p) → ¬p p + ... | no ¬p = no λ x → absurd (case subst Positive (sym q) x of λ (inc p) → ¬p p) Positive-pos : ∀ {x s p} → Positive (toℚ (possuc x / s [ p ])) Positive-pos = inc (pos _) @@ -221,19 +239,158 @@ operations. ```agda nonnegative-nonzero→positive : ∀ {x} → 0 ≤ x → x ≠ 0 → Positive x -nonnegative-nonzero→positive = work _ where - work : ∀ x → 0 ≤ x → x ≠ 0 → Positive x - work = ℚ-elim-propⁿ 1 λ where - d p posz (inc q) r → absurd (r (ext refl)) - d p (possuc x) (inc q) r → inc (pos x) - -*ℚ-positive : ∀ {x y} → Positive x → Positive y → Positive (x *ℚ y) -*ℚ-positive = work _ _ where - work : ∀ x y → Positive x → Positive y → Positive (x *ℚ y) - work = ℚ-elim-propⁿ 2 λ d p a b (inc x) (inc y) → inc (ℤ.*ℤ-positive x y) - -+ℚ-positive : ∀ {x y} → Positive x → Positive y → Positive (x +ℚ y) -+ℚ-positive = work _ _ where - work : ∀ x y → Positive x → Positive y → Positive (x +ℚ y) - work = ℚ-elim-propⁿ 2 λ d p a b (inc x) (inc y) → inc (ℤ.+ℤ-positive (ℤ.*ℤ-positive x p) (ℤ.*ℤ-positive y p)) +positive→nonzero : ∀ {x} → Positive x → x ≠ 0 +positive→nonnegative : ∀ {x} → Positive x → 0 ≤ x + +unquoteDef nonnegative-nonzero→positive positive→nonzero positive→nonnegative = do + by-elim-ℚ nonnegative-nonzero→positive λ where + d posz (inc q) r → absurd (r (ext refl)) + d (possuc x) (inc q) r → inc (pos x) + + by-elim-ℚ positive→nonzero λ where + d a (inc α) β → ℤ.positive→nonzero α (sym (ℤ.*ℤ-oner a) ∙ from-same-rational (unquotℚ β)) + + by-elim-ℚ positive→nonnegative λ where + d a (inc α) → inc (ℤ.≤-trans (ℤ.positive→nonnegative α) (ℤ.≤-refl' (sym (ℤ.*ℤ-oner a)))) +``` + +# The strict order on rationals + +With a bit more effort, we can also lift the strict ordering on the +integers to the rationals. The definitions end up being pretty much the +same as for lifting the partial order! We start by defining it on +fractions, then showing that it respects the quotient on both sides --- +we can't use transitivity and reflexivity like we did above, +unfortunately. + +```agda +_<ᶠ_ : Fraction → Fraction → Type +(x / s [ _ ]) <ᶠ (y / t [ _ ]) = (x *ℤ t) ℤ.< (y *ℤ s) + +<ᶠ-respr : ∀ {x y f} → x ≈ y → f <ᶠ y → f <ᶠ x +<ᶠ-respr {x / s [ p ]} {y / t [ q ]} {z / u [ r ]} xt~ys α = ℤ.*ℤ-cancel- +Having gotten the bulk of the construction out of the way, we +can lift it to a type family over pairs of rationals, and do the same +dance as to make this into something definitionally injective. + +```agda +private + ltℚ : ℚ → ℚ → Prop lzero + ltℚ (inc x) (inc y) = Coeq-rec₂ (hlevel 2) work (λ u (x , y , r) → r2 x y u r) (λ u (x , y , r) → r1 u x y r) x y where + work : Fraction → Fraction → Prop lzero + ∣ work x y ∣ = x <ᶠ y + work record{} record{} .is-tr = hlevel 1 + + r1 : ∀ u x y → x ≈ y → work u x ≡ work u y + r1 u@record{} x@record{} y@record{} p' = n-ua (prop-ext! + (λ α → <ᶠ-respr {y} {x} {u} (≈.symᶜ p') α) + (λ α → <ᶠ-respr {x} {y} {u} p' α)) + + r2 : ∀ x y u → x ≈ y → work x u ≡ work y u + r2 x@record{} y@record{} u@record{} p' = n-ua (prop-ext! + (λ α → <ᶠ-respl {x} {y} {u} p' α) + (λ α → <ᶠ-respl {y} {x} {u} (≈.symᶜ p') α)) + +record _<_ (x y : ℚ) : Type where + constructor inc + field + lower : ⌞ ltℚ x y ⌟ + +open _<_ +unquoteDecl H-Level-< = declare-record-hlevel 1 H-Level-< (quote _<_) +``` + + + +As before, everything we want to show about strict inequality on the +rationals follows by lifting the analogous facts from the integers, +where we're free to assume any set of related rationals has the same +denominator. + +```agda +instance + Dec-< : ∀ {x y} → Dec (x < y) + Dec-< {inc x} {inc y} = elim! {P = λ x → ∀ y → Dec (inc x < inc y)} go x y where + go : ∀ x y → Dec (toℚ x < toℚ y) + go x@record{} y@record{} with holds? ((↑ x *ℤ ↓ y) ℤ.< (↑ y *ℤ ↓ x)) + ... | yes p = yes (inc p) + ... | no ¬p = no (¬p ∘ lower) + +abstract + toℚ< : ∀ {x y} → (↑ x *ℤ ↓ y) ℤ.< (↑ y *ℤ ↓ x) → toℚ x < toℚ y + toℚ< {record{}} {record{}} p = inc p + + common-denom-< : ∀ {x y d} ⦃ _ : ℤ.Positive d ⦄ → x ℤ.< y → (x / d) < (y / d) + common-denom-< {x} {y} {d} p = inc (ℤ.*ℤ-preserves- + +```agda +module Data.Rational.Properties where +``` + +# Misc. properties of the rationals + + + +## Properties of multiplication + +```agda + *ℚ-distrib-minusr : ∀ {x y z} → (y -ℚ z) *ℚ x ≡ y *ℚ x -ℚ z *ℚ x + *ℚ-distrib-minusr {x@(inc _)} {y@(inc _)} {z@(inc _)} = *ℚ-distribr x y (-ℚ z) ∙ ap₂ _+ℚ_ (λ i → y *ℚ x) (ℚr.*-negatel {z} {x}) + + *ℚ-distrib-minusl : ∀ {x y z} → x *ℚ (y -ℚ z) ≡ x *ℚ y -ℚ x *ℚ z + *ℚ-distrib-minusl = *ℚ-commutative _ _ ∙ *ℚ-distrib-minusr ∙ ap₂ _-ℚ_ (*ℚ-commutative _ _) (*ℚ-commutative _ _) + + *ℚ-invl : ∀ {x} ⦃ p : Nonzero x ⦄ → invℚ x *ℚ x ≡ 1 + *ℚ-invl = *ℚ-commutative _ _ ∙ *ℚ-invr +``` + +## Nonzero rationals + +```agda + /ℚ-def : {x y : ℚ} ⦃ p : Nonzero y ⦄ → (x /ℚ y) ≡ x *ℚ invℚ y + /ℚ-def {inc x} {inc y} = refl + + *ℚ-nonzero : ∀ {x y} → Nonzero x → Nonzero y → Nonzero (x *ℚ y) + /ℚ-nonzero-num : ∀ {n d} ⦃ p : Nonzero d ⦄ → Nonzero (n /ℚ d) → Nonzero n + negℚ-nonzero : ∀ {n} → Nonzero n → Nonzero (-ℚ n) + invℚ-nonzero : ∀ {n} ⦃ d : Nonzero n ⦄ → Nonzero (invℚ n) + + unquoteDef *ℚ-nonzero /ℚ-nonzero-num negℚ-nonzero invℚ-nonzero = do + by-elim-ℚ *ℚ-nonzero λ d a b α β → + to-nonzero-frac (ℤ.*ℤ-nonzero (from-nonzero-frac α) (from-nonzero-frac β)) + + by-elim-ℚ /ℚ-nonzero-num λ where + d x ℤ.posz ⦃ p ⦄ nz → absurd (from-nonzero-frac p refl) + d x (ℤ.possuc y) nz → to-nonzero-frac (ℤ.*ℤ-nonzero-cancel {x} {d} (from-nonzero-frac nz)) + (ℤ.possuc x) ⦃ ℤ.pos .x ⦄ y (ℤ.negsuc z) nz → to-nonzero-frac (ℤ.*ℤ-nonzero-cancel {y} {ℤ.negsuc x} (from-nonzero-frac nz)) + + by-elim-ℚ negℚ-nonzero λ where + d ℤ.posz (inc α) → absurd (α (quotℚ (to-same-rational refl))) + d (ℤ.possuc x) (inc α) → to-nonzero-frac ℤ.negsuc≠pos + d (ℤ.negsuc x) (inc α) → to-nonzero-frac (suc≠zero ∘ ℤ.pos-injective) + + by-elim-ℚ invℚ-nonzero λ where + d ℤ.posz ⦃ inc α ⦄ → absurd (α (quotℚ (to-same-rational refl))) + (ℤ.possuc d) ⦃ ℤ.pos .d ⦄ (ℤ.possuc x) ⦃ inc α ⦄ → to-nonzero-frac (suc≠zero ∘ ℤ.pos-injective) + (ℤ.possuc d) ⦃ ℤ.pos .d ⦄ (ℤ.negsuc x) ⦃ inc α ⦄ → to-nonzero-frac ℤ.negsuc≠pos + + instance + Nonzero-*ℚ : ∀ {x y : ℚ} ⦃ p : Nonzero x ⦄ ⦃ q : Nonzero y ⦄ → Nonzero (x *ℚ y) + Nonzero-*ℚ ⦃ p ⦄ ⦃ q ⦄ = *ℚ-nonzero p q + + Nonzero-invℚ : ∀ {x} ⦃ p : Nonzero x ⦄ → Nonzero (invℚ x) + Nonzero-invℚ ⦃ p ⦄ = invℚ-nonzero ⦃ p ⦄ + + Nonzero-/ℚ : ∀ {x y} ⦃ p : Nonzero x ⦄ ⦃ q : Nonzero y ⦄ → Nonzero (x /ℚ y) + Nonzero-/ℚ {x} {y} ⦃ p ⦄ ⦃ q ⦄ = subst Nonzero (sym /ℚ-def) (*ℚ-nonzero p invℚ-nonzero) + + {-# OVERLAPPABLE Nonzero-*ℚ Nonzero-invℚ Nonzero-/ℚ #-} +``` + +## Properties of division + +```agda + /ℚ-oner : ∀ x → x /ℚ 1 ≡ x + /ℚ-oner (inc x) = *ℚ-idr (inc x) + + /ℚ-ap : ∀ {x y x' y'} {p : Nonzero y} {q : Nonzero y'} → x ≡ x' → y ≡ y' → (_/ℚ_ x y ⦃ p ⦄) ≡ (_/ℚ_ x' y' ⦃ q ⦄) + /ℚ-ap {p = α} {β} p q i = _/ℚ_ (p i) (q i) ⦃ is-prop→pathp (λ i → hlevel {T = Nonzero (q i)} 1) α β i ⦄ + + /ℚ-factorr : ∀ {x y} ⦃ p : Nonzero y ⦄ → (x *ℚ y) /ℚ y ≡ x + /ℚ-factorr = /ℚ-def ∙ sym (*ℚ-associative _ _ _) ∙ ap₂ _*ℚ_ refl *ℚ-invr ∙ *ℚ-idr _ + + /ℚ-self : ∀ {x} ⦃ p : Nonzero x ⦄ → x /ℚ x ≡ 1 + /ℚ-self {x} = /ℚ-def ∙ *ℚ-invr + + /ℚ-factorl : ∀ {x y} ⦃ p : Nonzero y ⦄ → (y *ℚ x) /ℚ y ≡ x + /ℚ-factorl = /ℚ-ap (*ℚ-commutative _ _) refl ∙ /ℚ-factorr + + /ℚ-scaler : ∀ {x y z} ⦃ p : Nonzero y ⦄ → (x /ℚ y) *ℚ z ≡ (x *ℚ z) /ℚ y + /ℚ-scaler = ap (_*ℚ _) /ℚ-def ∙ sym (*ℚ-associative _ _ _) ∙ ap (_ *ℚ_) (*ℚ-commutative _ _) ∙ *ℚ-associative _ _ _ ∙ sym /ℚ-def + + /ℚ-scalel : ∀ {x y z} ⦃ p : Nonzero z ⦄ → x *ℚ (y /ℚ z) ≡ (x *ℚ y) /ℚ z + /ℚ-scalel {z = z} = *ℚ-commutative _ _ ∙ /ℚ-scaler ∙ ap (λ e → e /ℚ z) (*ℚ-commutative _ _) + + /ℚ-cross : ∀ {q q' d} ⦃ α : Nonzero d ⦄ → q *ℚ d ≡ q' → q ≡ (q' /ℚ d) + /ℚ-cross {d = d} p = sym (ap (λ e → (e /ℚ d) ⦃ _ ⦄) (sym p) ∙ /ℚ-factorr) + + /ℚ-same + : ∀ {q q' d d'} ⦃ α : Nonzero d ⦄ ⦃ β : Nonzero d' ⦄ + → q *ℚ d' ≡ q' *ℚ d → q /ℚ d ≡ q' /ℚ d' + /ℚ-same p = /ℚ-cross (/ℚ-scaler ∙ sym (/ℚ-cross (sym p))) + + /ℚ-frac + : ∀ {n d} ⦃ p : ℤ.Positive d ⦄ + → (n / 1) /ℚ (d / 1) ≡ (n / d) + /ℚ-frac {n} {d = ℤ.possuc x} ⦃ p = ℤ.pos x ⦄ = quotℚ (to-same-rational (sym (ℤ.*ℤ-associative n 1 (ℤ.possuc x)))) + + invℚ-*ℚ + : ∀ {d d'} ⦃ p : Nonzero d ⦄ ⦃ p' : Nonzero d' ⦄ + → invℚ (d *ℚ d') ≡ invℚ d *ℚ invℚ d' + invℚ-*ℚ {d} {d'} = monoid-inverse-unique *ℚ-monoid (d *ℚ d') _ _ *ℚ-invl + (sym (*ℚ-associative _ _ _) ∙ ap (d *ℚ_) + (ap (d' *ℚ_) (*ℚ-commutative _ _) + ∙ *ℚ-associative d' _ _ + ∙ ap (_*ℚ invℚ d) *ℚ-invr ∙ *ℚ-idl (invℚ d)) + ∙ *ℚ-invr {d}) + + invℚ-invℚ : ∀ {d} ⦃ p : Nonzero d ⦄ → invℚ (invℚ d) ≡ d + invℚ-invℚ {d} = monoid-inverse-unique *ℚ-monoid (invℚ d) _ _ (*ℚ-commutative _ _ ∙ *ℚ-invr) (*ℚ-commutative _ _ ∙ *ℚ-invr) + + invℚ-/ℚ + : ∀ {q d} ⦃ p : Nonzero d ⦄ ⦃ p' : Nonzero q ⦄ + → invℚ (q /ℚ d) ≡ (d /ℚ q) + invℚ-/ℚ = ap₂ (λ e p → invℚ e ⦃ p ⦄) /ℚ-def prop! ∙ invℚ-*ℚ ∙ ap₂ _*ℚ_ refl invℚ-invℚ ∙ *ℚ-commutative _ _ ∙ sym /ℚ-def + + /ℚ-negatel + : ∀ {q d} ⦃ p : Nonzero d ⦄ → -ℚ (q /ℚ d) ≡ (-ℚ q) /ℚ d + /ℚ-negatel = ap -ℚ_ /ℚ-def ·· sym *ℚ-negatel ·· sym /ℚ-def + + /ℚ-def-+ℚ + : ∀ {q q' d d'} ⦃ p : Nonzero d ⦄ ⦃ p' : Nonzero d' ⦄ + → (q /ℚ d) +ℚ (q' /ℚ d') ≡ ((q *ℚ d' +ℚ q' *ℚ d) /ℚ (d *ℚ d')) + /ℚ-def-+ℚ {d = d} {d'} = /ℚ-cross + (*ℚ-distribr _ _ _ ∙ ap₂ _+ℚ_ + (/ℚ-scaler ∙ ap (λ e → e /ℚ d) (ap (_ *ℚ_) (*ℚ-commutative d d') ∙ *ℚ-associative _ d' d) ∙ /ℚ-factorr) + (/ℚ-scaler ∙ ap (λ e → e /ℚ d') (*ℚ-associative _ d d') ∙ /ℚ-factorr)) + + /ℚ-def-subℚ + : ∀ {q q' d d'} ⦃ p : Nonzero d ⦄ ⦃ p' : Nonzero d' ⦄ + → (q /ℚ d) -ℚ (q' /ℚ d') ≡ ((q *ℚ d' -ℚ q' *ℚ d) /ℚ (d *ℚ d')) + /ℚ-def-subℚ {d = d} {d'} = /ℚ-cross + (*ℚ-distrib-minusr ∙ ap₂ _-ℚ_ + (/ℚ-scaler ∙ ap (λ e → e /ℚ d) (ap (_ *ℚ_) (*ℚ-commutative d d') ∙ *ℚ-associative _ d' d) ∙ /ℚ-factorr) + (/ℚ-scaler ∙ ap (λ e → e /ℚ d') (*ℚ-associative _ d d') ∙ /ℚ-factorr)) + + /ℚ-def-*ℚ + : ∀ {q q' d d'} ⦃ p : Nonzero d ⦄ ⦃ p' : Nonzero d' ⦄ + → (q /ℚ d) *ℚ (q' /ℚ d') ≡ (q *ℚ q') /ℚ (d *ℚ d') + /ℚ-def-*ℚ {d = d} {d'} = /ℚ-cross + (sym (*ℚ-associative _ _ _) + ∙ ap (_ *ℚ_) (/ℚ-scaler ∙ ap (λ e → e /ℚ d') (*ℚ-associative _ d d') ∙ /ℚ-factorr) + ∙ /ℚ-scaler ∙ ap (λ e → e /ℚ d) (*ℚ-associative _ _ d) ∙ /ℚ-factorr) +``` + +## Positivity + +```agda +abstract + *ℚ-positive : ∀ {x y} → Positive x → Positive y → Positive (x *ℚ y) + +ℚ-positive : ∀ {x y} → Positive x → Positive y → Positive (x +ℚ y) + +ℚ-nonneg-positive : ∀ {x y} → 0 ≤ x → Positive y → Positive (x +ℚ y) + invℚ-positive : ∀ {d} (p : Positive d) → Positive (invℚ d ⦃ inc (positive→nonzero p) ⦄) + + unquoteDef *ℚ-positive +ℚ-positive invℚ-positive +ℚ-nonneg-positive = do + by-elim-ℚ *ℚ-positive λ d a b (inc x) (inc y) → inc (ℤ.*ℤ-positive x y) + by-elim-ℚ +ℚ-positive λ d a b (inc x) (inc y) → inc (ℤ.+ℤ-positive (ℤ.*ℤ-positive x auto) (ℤ.*ℤ-positive y auto)) + by-elim-ℚ invℚ-positive λ where + d@(ℤ.possuc d') ⦃ ℤ.pos .d' ⦄ (ℤ.possuc x) px → inc (ℤ.pos d') + + by-elim-ℚ +ℚ-nonneg-positive λ where + d (ℤ.posz) b (inc x) (inc y) → inc (subst ℤ.Positive (sym (ℤ.+ℤ-zerol (b ℤ.*ℤ d))) (ℤ.*ℤ-positive y auto)) + d (ℤ.possuc a) b (inc x) (inc y) → inc (ℤ.+ℤ-positive {ℤ.possuc a ℤ.*ℤ d} {b ℤ.*ℤ d} (ℤ.*ℤ-positive (ℤ.pos a) auto) (ℤ.*ℤ-positive y auto)) + + /ℚ-positive : ∀ {x y} (p : Positive x) (q : Positive y) → Positive ((x /ℚ y) ⦃ inc (positive→nonzero q) ⦄) + /ℚ-positive {y = y} p q = subst Positive + (sym (/ℚ-def ⦃ _ ⦄)) (*ℚ-positive p (invℚ-positive q)) + + from-positive : ∀ {x} → Positive x → 0 < x + to-positive : ∀ {x} → 0 < x → Positive x + + unquoteDef from-positive to-positive = do + by-elim-ℚ from-positive λ where + d (ℤ.possuc x) (inc (ℤ.pos .x)) → inc (ℤ.pos