From 6fc94a12ad2d59d40b8fa4dba45de6bb6e4beabb Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Mon, 1 May 2023 17:07:51 +0100 Subject: [PATCH 01/27] fresh commit of `FreeMagma` --- CHANGELOG.md | 6 + src/Algebra/Bundles/Free/Magma.agda | 443 ++++++++++++++++++++++++++++ 2 files changed, 449 insertions(+) create mode 100644 src/Algebra/Bundles/Free/Magma.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index 52b2a6762d..1ae0bb9728 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1370,6 +1370,12 @@ New modules Algebra.Construct.Flip.Op ``` +* Algebraic structures obtained as the free thing (for their signature): + ``` + Algebra.Bundles.Free + Algebra.Bundles.Free.Magma + ``` + * Morphisms between module-like algebraic structures: ``` Algebra.Module.Morphism.Construct.Composition diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda new file mode 100644 index 0000000000..45ff0fe06a --- /dev/null +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -0,0 +1,443 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Definitions of 'pre-free' and 'free' magma +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Algebra.Bundles.Free.Magma where + +open import Algebra.Core +open import Algebra.Bundles using (Magma) +open import Algebra.Bundles.Raw using (RawMagma) +import Algebra.Definitions as Definitions using (Congruent₂) +import Algebra.Structures as Structures using (IsMagma) +open import Algebra.Morphism.Structures using (IsMagmaHomomorphism) +import Algebra.Morphism.Construct.Identity as Identity +import Algebra.Morphism.Construct.Composition as Compose +open import Effect.Functor +open import Effect.Monad +open import Function.Base using (id; _∘_) +open import Relation.Binary.Core using (Rel) +open import Relation.Binary.Morphism using (IsRelHomomorphism) +open import Level using (Level; suc; _⊔_) +open import Relation.Nullary.Negation.Core using (¬_) +open import Relation.Binary + using (Setoid; IsEquivalence; Reflexive; Symmetric; Transitive) +open import Relation.Binary.Morphism.Bundles using (SetoidHomomorphism) +import Relation.Binary.Morphism.Construct.Identity as Identity +import Relation.Binary.Morphism.Construct.Composition as Compose +open import Relation.Binary.PropositionalEquality + using (_≡_; cong₂) renaming (refl to ≡-refl; isEquivalence to ≡-isEquivalence) +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning + +private + variable + ℓ a ℓa b ℓb c ℓc m ℓm n ℓn : Level + A : Set a + B : Set b + C : Set c + + +------------------------------------------------------------------------ +-- Morphisms between Magmas: belongs in its own place +-- Algebra.Morphism.Bundles +-- open import Algebra.Morphism.Bundles using (MagmaHomomorphism) +------------------------------------------------------------------------ + +record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a ⊔ b ⊔ ℓa ⊔ ℓb) where + + open Magma 𝓐 renaming (rawMagma to rawMagmaᴬ; setoid to setoidᴬ; Carrier to UA) + open Magma 𝓑 renaming (rawMagma to rawMagmaᴮ; setoid to setoidᴮ; Carrier to UB) + + field + ⟦_⟧ : UA → UB + + isMagmaHomomorphism : IsMagmaHomomorphism rawMagmaᴬ rawMagmaᴮ ⟦_⟧ + + open IsMagmaHomomorphism isMagmaHomomorphism public + + setoidHomomorphism : SetoidHomomorphism setoidᴬ setoidᴮ + setoidHomomorphism = record { ⟦_⟧ = ⟦_⟧ ; isRelHomomorphism = isRelHomomorphism } + +------------------------------------------------------------------------ +-- Syntax: 'pre'-free algebra + +module Syntax where + + infixl 7 _∙_ + + data Syntax (A : Set a) : Set a where + + var : A → Syntax A + _∙_ : Op₂ (Syntax A) + +-- Functor instance + + map : (A → B) → Syntax A → Syntax B + map f (var a) = var (f a) + map f (s ∙ t) = (map f s) ∙ (map f t) + + map-id : ∀ (t : Syntax A) → map id t ≡ t + map-id (var a) = ≡-refl + map-id (s ∙ t) = cong₂ _∙_ (map-id s) (map-id t) + + map-∘ : (g : A → B) → (f : B → C) → ∀ t → map (f ∘ g) t ≡ (map f ∘ map g) t + map-∘ g f (var a) = ≡-refl + map-∘ g f (s ∙ t) = cong₂ _∙_ (map-∘ g f s) (map-∘ g f t) + + syntaxRawFunctor : RawFunctor (Syntax {a}) + syntaxRawFunctor = record { _<$>_ = map } + +-- Monad instance + + bind : Syntax A → (A → Syntax B) → Syntax B + bind (var x) h = h x + bind (s ∙ t) h = bind s h ∙ bind t h + + syntaxRawMonad : RawMonad (Syntax {a}) + syntaxRawMonad = mkRawMonad Syntax var bind + + +------------------------------------------------------------------------ +-- parametrised 'equational' theory over the 'pre'-free algebra + +module EquationalTheory {A : Set a} (_≈ᴬ_ : Rel A ℓ) where + + open Syntax + + infix 4 _≈_ + + data _≈_ : Rel (Syntax A) (a ⊔ ℓ) + + open Definitions _≈_ + + data _≈_ where + + var : {a b : A} → a ≈ᴬ b → var a ≈ var b + _∙_ : Congruent₂ _∙_ + + refl : Reflexive _≈ᴬ_ → Reflexive _≈_ + refl r {var _} = var r + refl r {_ ∙ _} = (refl r) ∙ (refl r) + + sym : Symmetric _≈ᴬ_ → Symmetric _≈_ + sym s (var s₀) = var (s s₀) + sym s (s₁ ∙ s₂) = sym s s₁ ∙ sym s s₂ + + trans : Transitive _≈ᴬ_ → Transitive _≈_ + trans t (var r₀) (var s₀) = var (t r₀ s₀) + trans t (r₁ ∙ r₂) (s₁ ∙ s₂) = trans t r₁ s₁ ∙ trans t r₂ s₂ + + preservesEquivalence : IsEquivalence _≈ᴬ_ → IsEquivalence _≈_ + preservesEquivalence isEq = record + { refl = refl Eq.refl ; sym = sym Eq.sym ; trans = trans Eq.trans } + where module Eq = IsEquivalence isEq + + varIsRelHomomorphism : IsRelHomomorphism _≈ᴬ_ _≈_ var + varIsRelHomomorphism = record { cong = var } + + +------------------------------------------------------------------------ +-- Free algebra on a Set +{- + in the propositional case, we can immediately define the following + but how best to organise this under the Algebra.Bundles.Free hierarchy? + e.g. should we distinguish Free.Magma.Setoid from Free.Magma.Propositional? +-} + +module FreeRawMagma (A : Set a) where + + open Syntax + + open EquationalTheory {A = A} _≡_ + +-- inductively-defined equational theory conincides with _≡_ + + ≈⇒≡ : ∀ {m n} → m ≈ n → m ≡ n + ≈⇒≡ (var ≡-refl) = ≡-refl + ≈⇒≡ (eq₁ ∙ eq₂) = cong₂ _∙_ (≈⇒≡ eq₁) (≈⇒≡ eq₂) + + ≡⇒≈ : ∀ {m n} → m ≡ n → m ≈ n + ≡⇒≈ ≡-refl = refl ≡-refl + + freeRawMagma : RawMagma a a + freeRawMagma = record { Carrier = Syntax A ; _≈_ = _≡_ ; _∙_ = _∙_ } + + open Structures {A = Syntax A} _≡_ + + isMagma : IsMagma _∙_ + isMagma = record { isEquivalence = ≡-isEquivalence ; ∙-cong = cong₂ _∙_ } + + freeMagma : Magma a a + freeMagma = record { RawMagma freeRawMagma ; isMagma = isMagma } + + +------------------------------------------------------------------------ +-- Free algebra on a Setoid + +module FreeMagma (𝓐 : Setoid a ℓa) where + + open Setoid 𝓐 renaming (isEquivalence to isEqᴬ; _≈_ to _≈ᴬ_) + + open Syntax + + open EquationalTheory _≈ᴬ_ public + renaming (_≈_ to _≈ᵀ_) hiding (refl; sym; trans) + + open Structures _≈ᵀ_ + + isMagma : IsMagma _∙_ + isMagma = record { isEquivalence = preservesEquivalence isEqᴬ ; ∙-cong = _∙_ } + + freeMagma : Magma a (a ⊔ ℓa) + freeMagma = record { Carrier = Syntax Carrier; _≈_ = _≈ᵀ_ ; _∙_ = _∙_ ; isMagma = isMagma } + +-- re-export some substructure + + open Magma freeMagma public using (rawMagma; setoid; Carrier; _≈_) + + varSetoidHomomorphism : SetoidHomomorphism 𝓐 setoid + varSetoidHomomorphism = record { ⟦_⟧ = var; isRelHomomorphism = varIsRelHomomorphism } + + +------------------------------------------------------------------------ +-- Semantics: in terms of concrete Magma instances + +module _ (𝓜 : Magma m ℓm) where + + open Magma 𝓜 renaming (setoid to setoidᴹ; Carrier to UM; _∙_ to _∙ᴹ_) + open Syntax + +------------------------------------------------------------------------ +-- Eval, as the unique fold ⟦_⟧_ over Syntax + + module Eval (𝓐 : Setoid a ℓa) where + + open Setoid 𝓐 renaming (Carrier to UA) + + ⟦_⟧_ : Syntax UA → (UA → UM) → UM + ⟦ var a ⟧ η = η a + ⟦ s ∙ t ⟧ η = ⟦ s ⟧ η ∙ᴹ ⟦ t ⟧ η + +------------------------------------------------------------------------ +-- Any Magma *is* an algebra for the Syntax Functor + + alg : Syntax UM → UM + alg t = ⟦ t ⟧ id where open Eval setoidᴹ + +------------------------------------------------------------------------ +-- ⟦_⟧_ defines the (unique) lifting of Setoid homomorphisms to Magma homomorphisms + +module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) + (𝓗 : SetoidHomomorphism 𝓐 (Magma.setoid 𝓜)) where + + open Magma 𝓜 + renaming (Carrier to UM; _≈_ to _≈ᴹ_; _∙_ to _∙ᴹ_ + ; setoid to setoidᴹ; rawMagma to rawMagmaᴹ + ; isMagma to isMagmaᴹ) + + open ≈-Reasoning setoidᴹ + + open Syntax + + open Setoid 𝓐 renaming (Carrier to UA; _≈_ to _≈ᴬ_) + + open Eval 𝓜 𝓐 public + + open FreeMagma 𝓐 renaming (setoid to FA; Carrier to UFA) + + open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) + + private + + ⟦_⟧ᴹ : UFA → UM + ⟦_⟧ᴹ = ⟦_⟧ η + + open Structures _≈ᴹ_ + open IsMagma isMagmaᴹ renaming (∙-cong to congᴹ) + open IsRelHomomorphism hom-η renaming (cong to cong-η) + + module Existence where + + private + algᴹ = alg 𝓜 + + unfold-⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ᴹ ≈ᴹ algᴹ (map η t) + unfold-⟦ var a ⟧ᴹ = begin η a ∎ + unfold-⟦ s ∙ t ⟧ᴹ = congᴹ unfold-⟦ s ⟧ᴹ unfold-⟦ t ⟧ᴹ + + cong-⟦_⟧ : ∀ {s t} → s ≈ t → ⟦ s ⟧ᴹ ≈ᴹ ⟦ t ⟧ᴹ + cong-⟦ var r ⟧ = cong-η r + cong-⟦ s ∙ t ⟧ = congᴹ cong-⟦ s ⟧ cong-⟦ t ⟧ + + isRelHomomorphismᴹ : IsRelHomomorphism _≈_ _≈ᴹ_ ⟦_⟧ᴹ + isRelHomomorphismᴹ = record { cong = cong-⟦_⟧ } + + setoidHomomorphismᴹ : SetoidHomomorphism FA setoidᴹ + setoidHomomorphismᴹ = record { ⟦_⟧ = ⟦_⟧ᴹ ; isRelHomomorphism = isRelHomomorphismᴹ } + + isMagmaHomomorphismᴹ : IsMagmaHomomorphism rawMagma rawMagmaᴹ ⟦_⟧ᴹ + isMagmaHomomorphismᴹ = record { isRelHomomorphism = isRelHomomorphismᴹ + ; homo = λ s t → begin ⟦ s ⟧ᴹ ∙ᴹ ⟦ t ⟧ᴹ ∎ } + + magmaHomomorphismᴹ : MagmaHomomorphism freeMagma 𝓜 + magmaHomomorphismᴹ = record { ⟦_⟧ = ⟦_⟧ᴹ + ; isMagmaHomomorphism = isMagmaHomomorphismᴹ } + + record η-MagmaHomomorphism : Set (suc (a ⊔ m ⊔ ℓa ⊔ ℓm)) where + + field + magmaHomomorphism : MagmaHomomorphism freeMagma 𝓜 + open MagmaHomomorphism magmaHomomorphism public renaming (homo to ⟦⟧-homo) + field + ⟦_⟧∘var≈ᴹη : ∀ a → ⟦ var a ⟧ ≈ᴹ η a + + ⟦⟧ᴹ-η-MagmaHomomorphism : η-MagmaHomomorphism + ⟦⟧ᴹ-η-MagmaHomomorphism = record { magmaHomomorphism = Existence.magmaHomomorphismᴹ + ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var } + + module Uniqueness (η-magmaHomomorphism : η-MagmaHomomorphism) where + + open η-MagmaHomomorphism η-magmaHomomorphism + + isUnique⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ ≈ᴹ ⟦ t ⟧ᴹ + isUnique⟦ var a ⟧ᴹ = ⟦ a ⟧∘var≈ᴹη + isUnique⟦ s ∙ t ⟧ᴹ = begin + ⟦ s Syntax.∙ t ⟧ ≈⟨ ⟦⟧-homo s t ⟩ + ⟦ s ⟧ ∙ᴹ ⟦ t ⟧ ≈⟨ congᴹ isUnique⟦ s ⟧ᴹ isUnique⟦ t ⟧ᴹ ⟩ + ⟦ s ⟧ᴹ ∙ᴹ ⟦ t ⟧ᴹ ∎ + + module Corollary (𝓗 𝓚 : η-MagmaHomomorphism) + where + open η-MagmaHomomorphism 𝓗 renaming (⟦_⟧ to ⟦_⟧ᴴ) + open η-MagmaHomomorphism 𝓚 renaming (⟦_⟧ to ⟦_⟧ᴷ) + open Uniqueness 𝓗 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴴ) + open Uniqueness 𝓚 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴷ) + + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ ≈ᴹ ⟦ t ⟧ᴷ + isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧ᴹ ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ + +------------------------------------------------------------------------ +-- Immediate corollary: alg is in fact a MagmaHomomorphism + +module _ (𝓜 : Magma m ℓm) where + open Magma 𝓜 renaming (setoid to setoidᴹ; _≈_ to _≈ᴹ_; isMagma to isMagmaᴹ) + open FreeMagma setoidᴹ + + algMagmaHomomorphism : MagmaHomomorphism freeMagma 𝓜 + algMagmaHomomorphism = Existence.magmaHomomorphismᴹ + where open LeftAdjoint 𝓜 (Identity.setoidHomomorphism setoidᴹ) + + +------------------------------------------------------------------------ +-- Action of FreeMagma on Setoid homomorphisms + +module FreeMagmaFunctor {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} + (𝓗 : SetoidHomomorphism 𝓐 𝓑) where + + open FreeMagma 𝓐 renaming (freeMagma to freeMagmaᴬ) + open FreeMagma 𝓑 renaming (freeMagma to freeMagmaᴮ + ; varSetoidHomomorphism to 𝓥ᴮ) + + mapMagmaHomomorphism : MagmaHomomorphism freeMagmaᴬ freeMagmaᴮ + mapMagmaHomomorphism = Existence.magmaHomomorphismᴹ + where open LeftAdjoint freeMagmaᴮ (Compose.setoidHomomorphism 𝓗 𝓥ᴮ) + +------------------------------------------------------------------------ +-- Naturality of alg + +module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where + open Magma 𝓜 using () renaming (setoid to setoidᴹ) + open Magma 𝓝 using () renaming (_≈_ to _≈ᴺ_; refl to reflᴺ; trans to transᴺ) + + module _ (𝓕 : MagmaHomomorphism 𝓜 𝓝) where + open MagmaHomomorphism 𝓕 using (⟦_⟧; isMagmaHomomorphism; setoidHomomorphism) + open FreeMagmaFunctor setoidHomomorphism using (mapMagmaHomomorphism) + open MagmaHomomorphism mapMagmaHomomorphism + renaming (⟦_⟧ to map; isMagmaHomomorphism to map-isMagmaHomomorphism; setoidHomomorphism to mapSetoidHomomorphism) + open FreeMagma setoidᴹ renaming (freeMagma to freeMagmaᴹ) + open LeftAdjoint 𝓝 setoidHomomorphism + open MagmaHomomorphism (algMagmaHomomorphism 𝓜) + renaming (⟦_⟧ to algᴹ; isMagmaHomomorphism to algᴹ-isMagmaHomomorphism) + open MagmaHomomorphism (algMagmaHomomorphism 𝓝) + renaming (⟦_⟧ to algᴺ; isMagmaHomomorphism to algᴺ-isMagmaHomomorphism) + + naturality : ∀ t → ⟦ algᴹ t ⟧ ≈ᴺ algᴺ (map t) + naturality = Corollary.isUnique⟦_⟧ 𝓗 𝓚 + where + H K : MagmaHomomorphism freeMagmaᴹ 𝓝 + H = record { ⟦_⟧ = ⟦_⟧ ∘ algᴹ + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ algᴹ-isMagmaHomomorphism isMagmaHomomorphism } + + K = record { ⟦_⟧ = algᴺ ∘ map + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ map-isMagmaHomomorphism algᴺ-isMagmaHomomorphism } + + 𝓗 𝓚 : η-MagmaHomomorphism + 𝓗 = record { magmaHomomorphism = H ; ⟦_⟧∘var≈ᴹη = λ _ → reflᴺ } + 𝓚 = record { magmaHomomorphism = K ; ⟦_⟧∘var≈ᴹη = λ _ → reflᴺ } + + +------------------------------------------------------------------------ +-- Functoriality of FreeMagmaFunctor.map : by analogy with naturality + +module IdentityLaw (𝓐 : Setoid a ℓa) where + + open FreeMagma 𝓐 renaming (varSetoidHomomorphism to 𝓥) + open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) + + Id : MagmaHomomorphism freeMagma freeMagma + Id = record + { ⟦_⟧ = id + ; isMagmaHomomorphism = Identity.isMagmaHomomorphism rawMagma reflFA} + + open FreeMagmaFunctor (Identity.setoidHomomorphism 𝓐) + open MagmaHomomorphism mapMagmaHomomorphism renaming (⟦_⟧ to map-Id) + + map-id : ∀ t → map-Id t ≈FA t + map-id = Corollary.isUnique⟦_⟧ 𝓘ᴬ 𝓘 + where + open LeftAdjoint freeMagma 𝓥 + 𝓘ᴬ 𝓘 : η-MagmaHomomorphism + 𝓘ᴬ = record { magmaHomomorphism = mapMagmaHomomorphism ; ⟦_⟧∘var≈ᴹη = λ _ → reflFA } + 𝓘 = record { magmaHomomorphism = Id ; ⟦_⟧∘var≈ᴹη = λ _ → reflFA } + +module CompositionLaw + {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} {𝓒 : Setoid c ℓc} + (𝓗 : SetoidHomomorphism 𝓐 𝓑) (𝓚 : SetoidHomomorphism 𝓑 𝓒) where + + 𝓕 : SetoidHomomorphism 𝓐 𝓒 + 𝓕 = Compose.setoidHomomorphism 𝓗 𝓚 + + open FreeMagma 𝓐 renaming (freeMagma to freeMagmaA) + open FreeMagma 𝓑 renaming (freeMagma to freeMagmaB) + open FreeMagma 𝓒 renaming (freeMagma to freeMagmaC + ; setoid to setoidFC + ; varSetoidHomomorphism to 𝓥) + open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) + 𝓥∘𝓕 = Compose.setoidHomomorphism 𝓕 𝓥 + open FreeMagmaFunctor 𝓕 renaming (mapMagmaHomomorphism to MapAC) + open FreeMagmaFunctor 𝓗 renaming (mapMagmaHomomorphism to MapAB) + open FreeMagmaFunctor 𝓚 renaming (mapMagmaHomomorphism to MapBC) + open MagmaHomomorphism MapAC renaming (⟦_⟧ to mapAC) + open MagmaHomomorphism MapAB renaming (⟦_⟧ to mapAB; isMagmaHomomorphism to isMagmaAB) + open MagmaHomomorphism MapBC renaming (⟦_⟧ to mapBC; isMagmaHomomorphism to isMagmaBC) + + MapBC∘MapAB : MagmaHomomorphism freeMagmaA freeMagmaC + MapBC∘MapAB = record + { ⟦_⟧ = mapBC ∘ mapAB + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transFC isMagmaAB isMagmaBC} + + map-∘ : ∀ t → mapAC t ≈FC mapBC (mapAB t) + map-∘ = Corollary.isUnique⟦_⟧ 𝓐𝓒 𝓑𝓒∘𝓐𝓑 + where + open LeftAdjoint freeMagmaC 𝓥∘𝓕 + 𝓐𝓒 𝓑𝓒∘𝓐𝓑 : η-MagmaHomomorphism + 𝓐𝓒 = record { magmaHomomorphism = MapAC ; ⟦_⟧∘var≈ᴹη = λ _ → reflFC } + 𝓑𝓒∘𝓐𝓑 = record { magmaHomomorphism = MapBC∘MapAB ; ⟦_⟧∘var≈ᴹη = λ _ → reflFC } + + +------------------------------------------------------------------------ +-- Monad instance, etc.: TODO + From 36f0a9adf0593e862ffca0fd327fe286d98f52de Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Mon, 1 May 2023 17:09:31 +0100 Subject: [PATCH 02/27] fresh commit of `Free` --- src/Algebra/Bundles/Free.agda | 11 +++++++++++ 1 file changed, 11 insertions(+) create mode 100644 src/Algebra/Bundles/Free.agda diff --git a/src/Algebra/Bundles/Free.agda b/src/Algebra/Bundles/Free.agda new file mode 100644 index 0000000000..2a6771c5e5 --- /dev/null +++ b/src/Algebra/Bundles/Free.agda @@ -0,0 +1,11 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Definitions of 'pre-free' and 'free' bundles +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Algebra.Bundles.Free where + +open import Algebra.Bundles.Free.Magma public From 9cfe82375093fc3a70d87936e3ab8adf922c9c5e Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Fri, 12 May 2023 16:56:45 +0100 Subject: [PATCH 03/27] Update src/Algebra/Bundles/Free/Magma.agda Co-authored-by: G. Allais --- src/Algebra/Bundles/Free/Magma.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda index 45ff0fe06a..3109d000f6 100644 --- a/src/Algebra/Bundles/Free/Magma.agda +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -83,7 +83,7 @@ module Syntax where map-id (var a) = ≡-refl map-id (s ∙ t) = cong₂ _∙_ (map-id s) (map-id t) - map-∘ : (g : A → B) → (f : B → C) → ∀ t → map (f ∘ g) t ≡ (map f ∘ map g) t + map-∘ : (g : A → B) → (f : B → C) → map (f ∘ g) ≗ (map f ∘ map g) map-∘ g f (var a) = ≡-refl map-∘ g f (s ∙ t) = cong₂ _∙_ (map-∘ g f s) (map-∘ g f t) From 159cf04ac299ee218e23bbc111bc6d8f991033d6 Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Fri, 12 May 2023 16:57:06 +0100 Subject: [PATCH 04/27] Update src/Algebra/Bundles/Free/Magma.agda Co-authored-by: G. Allais --- src/Algebra/Bundles/Free/Magma.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda index 3109d000f6..c0db10379b 100644 --- a/src/Algebra/Bundles/Free/Magma.agda +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -79,7 +79,7 @@ module Syntax where map f (var a) = var (f a) map f (s ∙ t) = (map f s) ∙ (map f t) - map-id : ∀ (t : Syntax A) → map id t ≡ t + map-id : map {A = A} id ≗ id map-id (var a) = ≡-refl map-id (s ∙ t) = cong₂ _∙_ (map-id s) (map-id t) From 6f5381cd87b742ef612e1a63e7098ddea1990163 Mon Sep 17 00:00:00 2001 From: Guillaume Allais Date: Sat, 13 May 2023 10:07:42 +0100 Subject: [PATCH 05/27] [ fix ] scope error --- src/Algebra/Bundles/Free/Magma.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda index c0db10379b..fca4910ed2 100644 --- a/src/Algebra/Bundles/Free/Magma.agda +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -29,7 +29,7 @@ open import Relation.Binary.Morphism.Bundles using (SetoidHomomorphism) import Relation.Binary.Morphism.Construct.Identity as Identity import Relation.Binary.Morphism.Construct.Composition as Compose open import Relation.Binary.PropositionalEquality - using (_≡_; cong₂) renaming (refl to ≡-refl; isEquivalence to ≡-isEquivalence) + using (_≡_; _≗_; cong₂) renaming (refl to ≡-refl; isEquivalence to ≡-isEquivalence) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning private From 7dc217c8baad3698fee33af16fada242ea602f73 Mon Sep 17 00:00:00 2001 From: Guillaume Allais Date: Sat, 13 May 2023 10:07:54 +0100 Subject: [PATCH 06/27] [ fix ] whitespace violations --- src/Algebra/Bundles/Free/Magma.agda | 29 ++++++++++++++--------------- 1 file changed, 14 insertions(+), 15 deletions(-) diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda index fca4910ed2..2437f1e735 100644 --- a/src/Algebra/Bundles/Free/Magma.agda +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -60,7 +60,7 @@ record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a setoidHomomorphism : SetoidHomomorphism setoidᴬ setoidᴮ setoidHomomorphism = record { ⟦_⟧ = ⟦_⟧ ; isRelHomomorphism = isRelHomomorphism } - + ------------------------------------------------------------------------ -- Syntax: 'pre'-free algebra @@ -223,7 +223,7 @@ module _ (𝓜 : Magma m ℓm) where ------------------------------------------------------------------------ -- Any Magma *is* an algebra for the Syntax Functor - + alg : Syntax UM → UM alg t = ⟦ t ⟧ id where open Eval setoidᴹ @@ -248,19 +248,19 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) open FreeMagma 𝓐 renaming (setoid to FA; Carrier to UFA) - open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) + open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) private - + ⟦_⟧ᴹ : UFA → UM ⟦_⟧ᴹ = ⟦_⟧ η open Structures _≈ᴹ_ open IsMagma isMagmaᴹ renaming (∙-cong to congᴹ) open IsRelHomomorphism hom-η renaming (cong to cong-η) - + module Existence where - + private algᴹ = alg 𝓜 @@ -296,12 +296,12 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) ⟦⟧ᴹ-η-MagmaHomomorphism : η-MagmaHomomorphism ⟦⟧ᴹ-η-MagmaHomomorphism = record { magmaHomomorphism = Existence.magmaHomomorphismᴹ - ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var } - + ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var } + module Uniqueness (η-magmaHomomorphism : η-MagmaHomomorphism) where - + open η-MagmaHomomorphism η-magmaHomomorphism - + isUnique⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ ≈ᴹ ⟦ t ⟧ᴹ isUnique⟦ var a ⟧ᴹ = ⟦ a ⟧∘var≈ᴹη isUnique⟦ s ∙ t ⟧ᴹ = begin @@ -315,7 +315,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) open η-MagmaHomomorphism 𝓚 renaming (⟦_⟧ to ⟦_⟧ᴷ) open Uniqueness 𝓗 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴴ) open Uniqueness 𝓚 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴷ) - + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ ≈ᴹ ⟦ t ⟧ᴷ isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧ᴹ ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ @@ -325,7 +325,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) module _ (𝓜 : Magma m ℓm) where open Magma 𝓜 renaming (setoid to setoidᴹ; _≈_ to _≈ᴹ_; isMagma to isMagmaᴹ) open FreeMagma setoidᴹ - + algMagmaHomomorphism : MagmaHomomorphism freeMagma 𝓜 algMagmaHomomorphism = Existence.magmaHomomorphismᴹ where open LeftAdjoint 𝓜 (Identity.setoidHomomorphism setoidᴹ) @@ -385,7 +385,7 @@ module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where module IdentityLaw (𝓐 : Setoid a ℓa) where open FreeMagma 𝓐 renaming (varSetoidHomomorphism to 𝓥) - open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) + open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) Id : MagmaHomomorphism freeMagma freeMagma Id = record @@ -415,7 +415,7 @@ module CompositionLaw open FreeMagma 𝓒 renaming (freeMagma to freeMagmaC ; setoid to setoidFC ; varSetoidHomomorphism to 𝓥) - open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) + open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) 𝓥∘𝓕 = Compose.setoidHomomorphism 𝓕 𝓥 open FreeMagmaFunctor 𝓕 renaming (mapMagmaHomomorphism to MapAC) open FreeMagmaFunctor 𝓗 renaming (mapMagmaHomomorphism to MapAB) @@ -440,4 +440,3 @@ module CompositionLaw ------------------------------------------------------------------------ -- Monad instance, etc.: TODO - From dc9e6c5a0ecffa9fa4fa3ecbf91e9833affd23bc Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 14:41:14 +0100 Subject: [PATCH 07/27] review chnges --- src/Algebra/Bundles/Free/Magma.agda | 47 ++++++++++++++++------------- 1 file changed, 26 insertions(+), 21 deletions(-) diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda index 2437f1e735..0c2f08f443 100644 --- a/src/Algebra/Bundles/Free/Magma.agda +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -29,7 +29,8 @@ open import Relation.Binary.Morphism.Bundles using (SetoidHomomorphism) import Relation.Binary.Morphism.Construct.Identity as Identity import Relation.Binary.Morphism.Construct.Composition as Compose open import Relation.Binary.PropositionalEquality - using (_≡_; _≗_; cong₂) renaming (refl to ≡-refl; isEquivalence to ≡-isEquivalence) + using (_≡_; cong₂; _≗_) + renaming (refl to ≡-refl; isEquivalence to ≡-isEquivalence) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning private @@ -60,7 +61,7 @@ record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a setoidHomomorphism : SetoidHomomorphism setoidᴬ setoidᴮ setoidHomomorphism = record { ⟦_⟧ = ⟦_⟧ ; isRelHomomorphism = isRelHomomorphism } - + ------------------------------------------------------------------------ -- Syntax: 'pre'-free algebra @@ -130,9 +131,12 @@ module EquationalTheory {A : Set a} (_≈ᴬ_ : Rel A ℓ) where trans t (var r₀) (var s₀) = var (t r₀ s₀) trans t (r₁ ∙ r₂) (s₁ ∙ s₂) = trans t r₁ s₁ ∙ trans t r₂ s₂ - preservesEquivalence : IsEquivalence _≈ᴬ_ → IsEquivalence _≈_ - preservesEquivalence isEq = record - { refl = refl Eq.refl ; sym = sym Eq.sym ; trans = trans Eq.trans } + isEquivalence : IsEquivalence _≈ᴬ_ → IsEquivalence _≈_ + isEquivalence isEq = record + { refl = refl Eq.refl + ; sym = sym Eq.sym + ; trans = trans Eq.trans + } where module Eq = IsEquivalence isEq varIsRelHomomorphism : IsRelHomomorphism _≈ᴬ_ _≈_ var @@ -171,7 +175,7 @@ module FreeRawMagma (A : Set a) where isMagma = record { isEquivalence = ≡-isEquivalence ; ∙-cong = cong₂ _∙_ } freeMagma : Magma a a - freeMagma = record { RawMagma freeRawMagma ; isMagma = isMagma } + freeMagma = record { isMagma = isMagma } ------------------------------------------------------------------------ @@ -189,10 +193,10 @@ module FreeMagma (𝓐 : Setoid a ℓa) where open Structures _≈ᵀ_ isMagma : IsMagma _∙_ - isMagma = record { isEquivalence = preservesEquivalence isEqᴬ ; ∙-cong = _∙_ } + isMagma = record { isEquivalence = isEquivalence isEqᴬ ; ∙-cong = _∙_ } freeMagma : Magma a (a ⊔ ℓa) - freeMagma = record { Carrier = Syntax Carrier; _≈_ = _≈ᵀ_ ; _∙_ = _∙_ ; isMagma = isMagma } + freeMagma = record { isMagma = isMagma } -- re-export some substructure @@ -223,7 +227,7 @@ module _ (𝓜 : Magma m ℓm) where ------------------------------------------------------------------------ -- Any Magma *is* an algebra for the Syntax Functor - + alg : Syntax UM → UM alg t = ⟦ t ⟧ id where open Eval setoidᴹ @@ -248,19 +252,19 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) open FreeMagma 𝓐 renaming (setoid to FA; Carrier to UFA) - open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) + open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) private - + ⟦_⟧ᴹ : UFA → UM ⟦_⟧ᴹ = ⟦_⟧ η open Structures _≈ᴹ_ open IsMagma isMagmaᴹ renaming (∙-cong to congᴹ) open IsRelHomomorphism hom-η renaming (cong to cong-η) - + module Existence where - + private algᴹ = alg 𝓜 @@ -296,12 +300,12 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) ⟦⟧ᴹ-η-MagmaHomomorphism : η-MagmaHomomorphism ⟦⟧ᴹ-η-MagmaHomomorphism = record { magmaHomomorphism = Existence.magmaHomomorphismᴹ - ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var } - + ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var } + module Uniqueness (η-magmaHomomorphism : η-MagmaHomomorphism) where - + open η-MagmaHomomorphism η-magmaHomomorphism - + isUnique⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ ≈ᴹ ⟦ t ⟧ᴹ isUnique⟦ var a ⟧ᴹ = ⟦ a ⟧∘var≈ᴹη isUnique⟦ s ∙ t ⟧ᴹ = begin @@ -315,7 +319,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) open η-MagmaHomomorphism 𝓚 renaming (⟦_⟧ to ⟦_⟧ᴷ) open Uniqueness 𝓗 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴴ) open Uniqueness 𝓚 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴷ) - + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ ≈ᴹ ⟦ t ⟧ᴷ isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧ᴹ ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ @@ -325,7 +329,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) module _ (𝓜 : Magma m ℓm) where open Magma 𝓜 renaming (setoid to setoidᴹ; _≈_ to _≈ᴹ_; isMagma to isMagmaᴹ) open FreeMagma setoidᴹ - + algMagmaHomomorphism : MagmaHomomorphism freeMagma 𝓜 algMagmaHomomorphism = Existence.magmaHomomorphismᴹ where open LeftAdjoint 𝓜 (Identity.setoidHomomorphism setoidᴹ) @@ -385,7 +389,7 @@ module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where module IdentityLaw (𝓐 : Setoid a ℓa) where open FreeMagma 𝓐 renaming (varSetoidHomomorphism to 𝓥) - open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) + open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) Id : MagmaHomomorphism freeMagma freeMagma Id = record @@ -415,7 +419,7 @@ module CompositionLaw open FreeMagma 𝓒 renaming (freeMagma to freeMagmaC ; setoid to setoidFC ; varSetoidHomomorphism to 𝓥) - open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) + open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) 𝓥∘𝓕 = Compose.setoidHomomorphism 𝓕 𝓥 open FreeMagmaFunctor 𝓕 renaming (mapMagmaHomomorphism to MapAC) open FreeMagmaFunctor 𝓗 renaming (mapMagmaHomomorphism to MapAB) @@ -440,3 +444,4 @@ module CompositionLaw ------------------------------------------------------------------------ -- Monad instance, etc.: TODO + From 4dd32110327862e935d9d1cde0ee411c36d302cf Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 14:53:32 +0100 Subject: [PATCH 08/27] more review chnges --- src/Algebra/Bundles/Free/Magma.agda | 41 ++++++++++++++++++++--------- 1 file changed, 28 insertions(+), 13 deletions(-) diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda index 0c2f08f443..398f40ec97 100644 --- a/src/Algebra/Bundles/Free/Magma.agda +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -74,6 +74,9 @@ module Syntax where var : A → Syntax A _∙_ : Op₂ (Syntax A) + _∙-cong_ : ∀ {s s′ t t′ : Syntax A} → s ≡ s′ → t ≡ t′ → s ∙ t ≡ s′ ∙ t′ + _∙-cong_ = cong₂ _∙_ + -- Functor instance map : (A → B) → Syntax A → Syntax B @@ -82,11 +85,11 @@ module Syntax where map-id : map {A = A} id ≗ id map-id (var a) = ≡-refl - map-id (s ∙ t) = cong₂ _∙_ (map-id s) (map-id t) + map-id (s ∙ t) = (map-id s) ∙-cong (map-id t) map-∘ : (g : A → B) → (f : B → C) → map (f ∘ g) ≗ (map f ∘ map g) map-∘ g f (var a) = ≡-refl - map-∘ g f (s ∙ t) = cong₂ _∙_ (map-∘ g f s) (map-∘ g f t) + map-∘ g f (s ∙ t) = (map-∘ g f s) ∙-cong (map-∘ g f t) syntaxRawFunctor : RawFunctor (Syntax {a}) syntaxRawFunctor = record { _<$>_ = map } @@ -161,7 +164,7 @@ module FreeRawMagma (A : Set a) where ≈⇒≡ : ∀ {m n} → m ≈ n → m ≡ n ≈⇒≡ (var ≡-refl) = ≡-refl - ≈⇒≡ (eq₁ ∙ eq₂) = cong₂ _∙_ (≈⇒≡ eq₁) (≈⇒≡ eq₂) + ≈⇒≡ (eq₁ ∙ eq₂) = (≈⇒≡ eq₁) ∙-cong (≈⇒≡ eq₂) ≡⇒≈ : ∀ {m n} → m ≡ n → m ≈ n ≡⇒≈ ≡-refl = refl ≡-refl @@ -172,7 +175,7 @@ module FreeRawMagma (A : Set a) where open Structures {A = Syntax A} _≡_ isMagma : IsMagma _∙_ - isMagma = record { isEquivalence = ≡-isEquivalence ; ∙-cong = cong₂ _∙_ } + isMagma = record { isEquivalence = ≡-isEquivalence ; ∙-cong = _∙-cong_ } freeMagma : Magma a a freeMagma = record { isMagma = isMagma } @@ -193,7 +196,10 @@ module FreeMagma (𝓐 : Setoid a ℓa) where open Structures _≈ᵀ_ isMagma : IsMagma _∙_ - isMagma = record { isEquivalence = isEquivalence isEqᴬ ; ∙-cong = _∙_ } + isMagma = record + { isEquivalence = isEquivalence isEqᴬ + ; ∙-cong = _∙_ + } freeMagma : Magma a (a ⊔ ℓa) freeMagma = record { isMagma = isMagma } @@ -203,7 +209,10 @@ module FreeMagma (𝓐 : Setoid a ℓa) where open Magma freeMagma public using (rawMagma; setoid; Carrier; _≈_) varSetoidHomomorphism : SetoidHomomorphism 𝓐 setoid - varSetoidHomomorphism = record { ⟦_⟧ = var; isRelHomomorphism = varIsRelHomomorphism } + varSetoidHomomorphism = record + { ⟦_⟧ = var + ; isRelHomomorphism = varIsRelHomomorphism + } ------------------------------------------------------------------------ @@ -283,8 +292,10 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) setoidHomomorphismᴹ = record { ⟦_⟧ = ⟦_⟧ᴹ ; isRelHomomorphism = isRelHomomorphismᴹ } isMagmaHomomorphismᴹ : IsMagmaHomomorphism rawMagma rawMagmaᴹ ⟦_⟧ᴹ - isMagmaHomomorphismᴹ = record { isRelHomomorphism = isRelHomomorphismᴹ - ; homo = λ s t → begin ⟦ s ⟧ᴹ ∙ᴹ ⟦ t ⟧ᴹ ∎ } + isMagmaHomomorphismᴹ = record + { isRelHomomorphism = isRelHomomorphismᴹ + ; homo = λ s t → begin ⟦ s ⟧ᴹ ∙ᴹ ⟦ t ⟧ᴹ ∎ + } magmaHomomorphismᴹ : MagmaHomomorphism freeMagma 𝓜 magmaHomomorphismᴹ = record { ⟦_⟧ = ⟦_⟧ᴹ @@ -299,8 +310,10 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) ⟦_⟧∘var≈ᴹη : ∀ a → ⟦ var a ⟧ ≈ᴹ η a ⟦⟧ᴹ-η-MagmaHomomorphism : η-MagmaHomomorphism - ⟦⟧ᴹ-η-MagmaHomomorphism = record { magmaHomomorphism = Existence.magmaHomomorphismᴹ - ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var } + ⟦⟧ᴹ-η-MagmaHomomorphism = record + { magmaHomomorphism = Existence.magmaHomomorphismᴹ + ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var + } module Uniqueness (η-magmaHomomorphism : η-MagmaHomomorphism) where @@ -372,7 +385,8 @@ module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where naturality = Corollary.isUnique⟦_⟧ 𝓗 𝓚 where H K : MagmaHomomorphism freeMagmaᴹ 𝓝 - H = record { ⟦_⟧ = ⟦_⟧ ∘ algᴹ + H = record + { ⟦_⟧ = ⟦_⟧ ∘ algᴹ ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ algᴹ-isMagmaHomomorphism isMagmaHomomorphism } K = record { ⟦_⟧ = algᴺ ∘ map @@ -393,8 +407,9 @@ module IdentityLaw (𝓐 : Setoid a ℓa) where Id : MagmaHomomorphism freeMagma freeMagma Id = record - { ⟦_⟧ = id - ; isMagmaHomomorphism = Identity.isMagmaHomomorphism rawMagma reflFA} + { ⟦_⟧ = id + ; isMagmaHomomorphism = Identity.isMagmaHomomorphism rawMagma reflFA + } open FreeMagmaFunctor (Identity.setoidHomomorphism 𝓐) open MagmaHomomorphism mapMagmaHomomorphism renaming (⟦_⟧ to map-Id) From edcafaafca5d1bd01aa3a5b296569987878585a7 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 14:55:02 +0100 Subject: [PATCH 09/27] fix whitespace --- src/Algebra/Bundles/Free/Magma.agda | 28 ++++++++++++++-------------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Bundles/Free/Magma.agda index 398f40ec97..cd8dd72601 100644 --- a/src/Algebra/Bundles/Free/Magma.agda +++ b/src/Algebra/Bundles/Free/Magma.agda @@ -61,7 +61,7 @@ record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a setoidHomomorphism : SetoidHomomorphism setoidᴬ setoidᴮ setoidHomomorphism = record { ⟦_⟧ = ⟦_⟧ ; isRelHomomorphism = isRelHomomorphism } - + ------------------------------------------------------------------------ -- Syntax: 'pre'-free algebra @@ -236,7 +236,7 @@ module _ (𝓜 : Magma m ℓm) where ------------------------------------------------------------------------ -- Any Magma *is* an algebra for the Syntax Functor - + alg : Syntax UM → UM alg t = ⟦ t ⟧ id where open Eval setoidᴹ @@ -261,19 +261,19 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) open FreeMagma 𝓐 renaming (setoid to FA; Carrier to UFA) - open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) + open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) private - + ⟦_⟧ᴹ : UFA → UM ⟦_⟧ᴹ = ⟦_⟧ η open Structures _≈ᴹ_ open IsMagma isMagmaᴹ renaming (∙-cong to congᴹ) open IsRelHomomorphism hom-η renaming (cong to cong-η) - + module Existence where - + private algᴹ = alg 𝓜 @@ -313,12 +313,12 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) ⟦⟧ᴹ-η-MagmaHomomorphism = record { magmaHomomorphism = Existence.magmaHomomorphismᴹ ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var - } - + } + module Uniqueness (η-magmaHomomorphism : η-MagmaHomomorphism) where - + open η-MagmaHomomorphism η-magmaHomomorphism - + isUnique⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ ≈ᴹ ⟦ t ⟧ᴹ isUnique⟦ var a ⟧ᴹ = ⟦ a ⟧∘var≈ᴹη isUnique⟦ s ∙ t ⟧ᴹ = begin @@ -332,7 +332,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) open η-MagmaHomomorphism 𝓚 renaming (⟦_⟧ to ⟦_⟧ᴷ) open Uniqueness 𝓗 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴴ) open Uniqueness 𝓚 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴷ) - + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ ≈ᴹ ⟦ t ⟧ᴷ isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧ᴹ ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ @@ -342,7 +342,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) module _ (𝓜 : Magma m ℓm) where open Magma 𝓜 renaming (setoid to setoidᴹ; _≈_ to _≈ᴹ_; isMagma to isMagmaᴹ) open FreeMagma setoidᴹ - + algMagmaHomomorphism : MagmaHomomorphism freeMagma 𝓜 algMagmaHomomorphism = Existence.magmaHomomorphismᴹ where open LeftAdjoint 𝓜 (Identity.setoidHomomorphism setoidᴹ) @@ -403,7 +403,7 @@ module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where module IdentityLaw (𝓐 : Setoid a ℓa) where open FreeMagma 𝓐 renaming (varSetoidHomomorphism to 𝓥) - open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) + open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) Id : MagmaHomomorphism freeMagma freeMagma Id = record @@ -434,7 +434,7 @@ module CompositionLaw open FreeMagma 𝓒 renaming (freeMagma to freeMagmaC ; setoid to setoidFC ; varSetoidHomomorphism to 𝓥) - open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) + open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) 𝓥∘𝓕 = Compose.setoidHomomorphism 𝓕 𝓥 open FreeMagmaFunctor 𝓕 renaming (mapMagmaHomomorphism to MapAC) open FreeMagmaFunctor 𝓗 renaming (mapMagmaHomomorphism to MapAB) From 1f98554dddd02de271141a0662ce25e666d7a955 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 14:57:19 +0100 Subject: [PATCH 10/27] moved to `Algebra.Free.X` --- src/Algebra/{Bundles => }/Free.agda | 0 src/Algebra/{Bundles => }/Free/Magma.agda | 0 2 files changed, 0 insertions(+), 0 deletions(-) rename src/Algebra/{Bundles => }/Free.agda (100%) rename src/Algebra/{Bundles => }/Free/Magma.agda (100%) diff --git a/src/Algebra/Bundles/Free.agda b/src/Algebra/Free.agda similarity index 100% rename from src/Algebra/Bundles/Free.agda rename to src/Algebra/Free.agda diff --git a/src/Algebra/Bundles/Free/Magma.agda b/src/Algebra/Free/Magma.agda similarity index 100% rename from src/Algebra/Bundles/Free/Magma.agda rename to src/Algebra/Free/Magma.agda From 652d4b7289dd4e4ff095f49a46b3808bc1f0e81e Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 15:05:47 +0100 Subject: [PATCH 11/27] more spacing of records --- src/Algebra/Free/Magma.agda | 24 ++++++++++-------------- 1 file changed, 10 insertions(+), 14 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index cd8dd72601..7bb69b9d8a 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -6,7 +6,7 @@ {-# OPTIONS --cubical-compatible --safe #-} -module Algebra.Bundles.Free.Magma where +module Algebra.Free.Magma where open import Algebra.Core open import Algebra.Bundles using (Magma) @@ -113,14 +113,10 @@ module EquationalTheory {A : Set a} (_≈ᴬ_ : Rel A ℓ) where infix 4 _≈_ - data _≈_ : Rel (Syntax A) (a ⊔ ℓ) - - open Definitions _≈_ - - data _≈_ where + data _≈_ : Rel (Syntax A) (a ⊔ ℓ) where var : {a b : A} → a ≈ᴬ b → var a ≈ var b - _∙_ : Congruent₂ _∙_ + _∙_ : Definitions.Congruent₂ _≈_ _∙_ refl : Reflexive _≈ᴬ_ → Reflexive _≈_ refl r {var _} = var r @@ -274,10 +270,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) module Existence where - private - algᴹ = alg 𝓜 - - unfold-⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ᴹ ≈ᴹ algᴹ (map η t) + unfold-⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ᴹ ≈ᴹ alg 𝓜 (map η t) unfold-⟦ var a ⟧ᴹ = begin η a ∎ unfold-⟦ s ∙ t ⟧ᴹ = congᴹ unfold-⟦ s ⟧ᴹ unfold-⟦ t ⟧ᴹ @@ -389,8 +382,10 @@ module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where { ⟦_⟧ = ⟦_⟧ ∘ algᴹ ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ algᴹ-isMagmaHomomorphism isMagmaHomomorphism } - K = record { ⟦_⟧ = algᴺ ∘ map - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ map-isMagmaHomomorphism algᴺ-isMagmaHomomorphism } + K = record + { ⟦_⟧ = algᴺ ∘ map + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ map-isMagmaHomomorphism algᴺ-isMagmaHomomorphism + } 𝓗 𝓚 : η-MagmaHomomorphism 𝓗 = record { magmaHomomorphism = H ; ⟦_⟧∘var≈ᴹη = λ _ → reflᴺ } @@ -446,7 +441,8 @@ module CompositionLaw MapBC∘MapAB : MagmaHomomorphism freeMagmaA freeMagmaC MapBC∘MapAB = record { ⟦_⟧ = mapBC ∘ mapAB - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transFC isMagmaAB isMagmaBC} + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transFC isMagmaAB isMagmaBC + } map-∘ : ∀ t → mapAC t ≈FC mapBC (mapAB t) map-∘ = Corollary.isUnique⟦_⟧ 𝓐𝓒 𝓑𝓒∘𝓐𝓑 From f37c57a55b7c43bfb1666e57f1ffdcd1882e5835 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 15:26:59 +0100 Subject: [PATCH 12/27] refactored to separate out bundled `MagmaHomomorphism` --- src/Algebra/Free/Magma.agda | 12 +++++----- src/Algebra/Morphism/Bundles.agda | 39 +++++++++++++++++++++++++++++++ 2 files changed, 45 insertions(+), 6 deletions(-) create mode 100644 src/Algebra/Morphism/Bundles.agda diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 7bb69b9d8a..9cbe50f52e 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -13,6 +13,7 @@ open import Algebra.Bundles using (Magma) open import Algebra.Bundles.Raw using (RawMagma) import Algebra.Definitions as Definitions using (Congruent₂) import Algebra.Structures as Structures using (IsMagma) +open import Algebra.Morphism.Bundles using (MagmaHomomorphism) open import Algebra.Morphism.Structures using (IsMagmaHomomorphism) import Algebra.Morphism.Construct.Identity as Identity import Algebra.Morphism.Construct.Composition as Compose @@ -48,18 +49,17 @@ private ------------------------------------------------------------------------ record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a ⊔ b ⊔ ℓa ⊔ ℓb) where - - open Magma 𝓐 renaming (rawMagma to rawMagmaᴬ; setoid to setoidᴬ; Carrier to UA) - open Magma 𝓑 renaming (rawMagma to rawMagmaᴮ; setoid to setoidᴮ; Carrier to UB) + module A = Magma 𝓐 + module B = Magma 𝓑 field - ⟦_⟧ : UA → UB + ⟦_⟧ : A.Carrier → B.Carrier - isMagmaHomomorphism : IsMagmaHomomorphism rawMagmaᴬ rawMagmaᴮ ⟦_⟧ + isMagmaHomomorphism : IsMagmaHomomorphism A.rawMagma B.rawMagma ⟦_⟧ open IsMagmaHomomorphism isMagmaHomomorphism public - setoidHomomorphism : SetoidHomomorphism setoidᴬ setoidᴮ + setoidHomomorphism : SetoidHomomorphism A.setoid B.setoid setoidHomomorphism = record { ⟦_⟧ = ⟦_⟧ ; isRelHomomorphism = isRelHomomorphism } ------------------------------------------------------------------------ diff --git a/src/Algebra/Morphism/Bundles.agda b/src/Algebra/Morphism/Bundles.agda new file mode 100644 index 0000000000..527443bbde --- /dev/null +++ b/src/Algebra/Morphism/Bundles.agda @@ -0,0 +1,39 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Bundled definitions of homomorphisms between Magmas +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Algebra.Morphism.Bundles where + +open import Algebra.Bundles using (Magma) +open import Algebra.Morphism.Structures using (IsMagmaHomomorphism) +open import Level using (Level; suc; _⊔_) +open import Relation.Binary.Morphism using (IsRelHomomorphism) +open import Relation.Binary.Morphism.Bundles using (SetoidHomomorphism) + +private + variable + ℓ a ℓa b ℓb : Level + + +------------------------------------------------------------------------ +-- Morphisms between Magmas +------------------------------------------------------------------------ + +record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a ⊔ b ⊔ ℓa ⊔ ℓb) where + module A = Magma 𝓐 + module B = Magma 𝓑 + + field + ⟦_⟧ : A.Carrier → B.Carrier + + isMagmaHomomorphism : IsMagmaHomomorphism A.rawMagma B.rawMagma ⟦_⟧ + + open IsMagmaHomomorphism isMagmaHomomorphism public + + setoidHomomorphism : SetoidHomomorphism A.setoid B.setoid + setoidHomomorphism = record { ⟦_⟧ = ⟦_⟧ ; isRelHomomorphism = isRelHomomorphism } + From 8d3fc8e2f717929cab015fd8a5334ad5ec6ee3d8 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 15:36:46 +0100 Subject: [PATCH 13/27] corrected imports; `CHANGELOG` --- CHANGELOG.md | 9 +++-- src/Algebra/Free.agda | 4 +-- src/Algebra/Free/Magma.agda | 65 +++++++++++++++---------------------- 3 files changed, 35 insertions(+), 43 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 1ae0bb9728..e010795900 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1372,8 +1372,8 @@ New modules * Algebraic structures obtained as the free thing (for their signature): ``` - Algebra.Bundles.Free - Algebra.Bundles.Free.Magma + Algebra.Free + Algebra.Free.Magma ``` * Morphisms between module-like algebraic structures: @@ -1385,6 +1385,11 @@ New modules Algebra.Module.Properties ``` +* Bundled morphisms between algebraic structures: + ``` + Algebra.Morphism.Bundles + ``` + * Identity morphisms and composition of morphisms between algebraic structures: ``` Algebra.Morphism.Construct.Composition diff --git a/src/Algebra/Free.agda b/src/Algebra/Free.agda index 2a6771c5e5..a2f1f331bd 100644 --- a/src/Algebra/Free.agda +++ b/src/Algebra/Free.agda @@ -6,6 +6,6 @@ {-# OPTIONS --cubical-compatible --safe #-} -module Algebra.Bundles.Free where +module Algebra.Free where -open import Algebra.Bundles.Free.Magma public +open import Algebra.Free.Magma public diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 9cbe50f52e..625e77370e 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -42,26 +42,6 @@ private C : Set c ------------------------------------------------------------------------- --- Morphisms between Magmas: belongs in its own place --- Algebra.Morphism.Bundles --- open import Algebra.Morphism.Bundles using (MagmaHomomorphism) ------------------------------------------------------------------------- - -record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a ⊔ b ⊔ ℓa ⊔ ℓb) where - module A = Magma 𝓐 - module B = Magma 𝓑 - - field - ⟦_⟧ : A.Carrier → B.Carrier - - isMagmaHomomorphism : IsMagmaHomomorphism A.rawMagma B.rawMagma ⟦_⟧ - - open IsMagmaHomomorphism isMagmaHomomorphism public - - setoidHomomorphism : SetoidHomomorphism A.setoid B.setoid - setoidHomomorphism = record { ⟦_⟧ = ⟦_⟧ ; isRelHomomorphism = isRelHomomorphism } - ------------------------------------------------------------------------ -- Syntax: 'pre'-free algebra @@ -166,12 +146,19 @@ module FreeRawMagma (A : Set a) where ≡⇒≈ ≡-refl = refl ≡-refl freeRawMagma : RawMagma a a - freeRawMagma = record { Carrier = Syntax A ; _≈_ = _≡_ ; _∙_ = _∙_ } + freeRawMagma = record + { Carrier = Syntax A + ; _≈_ = _≡_ + ; _∙_ = _∙_ + } open Structures {A = Syntax A} _≡_ isMagma : IsMagma _∙_ - isMagma = record { isEquivalence = ≡-isEquivalence ; ∙-cong = _∙-cong_ } + isMagma = record + { isEquivalence = ≡-isEquivalence + ; ∙-cong = _∙-cong_ + } freeMagma : Magma a a freeMagma = record { isMagma = isMagma } @@ -182,33 +169,33 @@ module FreeRawMagma (A : Set a) where module FreeMagma (𝓐 : Setoid a ℓa) where - open Setoid 𝓐 renaming (isEquivalence to isEqᴬ; _≈_ to _≈ᴬ_) + module A = Setoid 𝓐 open Syntax - open EquationalTheory _≈ᴬ_ public - renaming (_≈_ to _≈ᵀ_) hiding (refl; sym; trans) + open EquationalTheory A._≈_ public + hiding (refl; sym; trans) - open Structures _≈ᵀ_ + open Structures _≈_ isMagma : IsMagma _∙_ isMagma = record - { isEquivalence = isEquivalence isEqᴬ - ; ∙-cong = _∙_ - } + { isEquivalence = isEquivalence A.isEquivalence + ; ∙-cong = _∙_ + } freeMagma : Magma a (a ⊔ ℓa) freeMagma = record { isMagma = isMagma } -- re-export some substructure - open Magma freeMagma public using (rawMagma; setoid; Carrier; _≈_) + open Magma freeMagma public using (rawMagma; setoid; Carrier) varSetoidHomomorphism : SetoidHomomorphism 𝓐 setoid varSetoidHomomorphism = record - { ⟦_⟧ = var - ; isRelHomomorphism = varIsRelHomomorphism - } + { ⟦_⟧ = var + ; isRelHomomorphism = varIsRelHomomorphism + } ------------------------------------------------------------------------ @@ -216,7 +203,7 @@ module FreeMagma (𝓐 : Setoid a ℓa) where module _ (𝓜 : Magma m ℓm) where - open Magma 𝓜 renaming (setoid to setoidᴹ; Carrier to UM; _∙_ to _∙ᴹ_) + module M = Magma 𝓜 open Syntax ------------------------------------------------------------------------ @@ -224,17 +211,17 @@ module _ (𝓜 : Magma m ℓm) where module Eval (𝓐 : Setoid a ℓa) where - open Setoid 𝓐 renaming (Carrier to UA) + module A = Setoid 𝓐 - ⟦_⟧_ : Syntax UA → (UA → UM) → UM + ⟦_⟧_ : Syntax A.Carrier → (A.Carrier → M.Carrier) → M.Carrier ⟦ var a ⟧ η = η a - ⟦ s ∙ t ⟧ η = ⟦ s ⟧ η ∙ᴹ ⟦ t ⟧ η + ⟦ s ∙ t ⟧ η = ⟦ s ⟧ η M.∙ ⟦ t ⟧ η ------------------------------------------------------------------------ -- Any Magma *is* an algebra for the Syntax Functor - alg : Syntax UM → UM - alg t = ⟦ t ⟧ id where open Eval setoidᴹ + alg : Syntax M.Carrier → M.Carrier + alg t = ⟦ t ⟧ id where open Eval M.setoid ------------------------------------------------------------------------ -- ⟦_⟧_ defines the (unique) lifting of Setoid homomorphisms to Magma homomorphisms From 111229e81a82229ec6b43450b62025e315919897 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 16:38:11 +0100 Subject: [PATCH 14/27] systematically refactored to use `private` modules --- src/Algebra/Free/Magma.agda | 205 +++++++++++++++--------------- src/Algebra/Morphism/Bundles.agda | 5 +- 2 files changed, 108 insertions(+), 102 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 625e77370e..4c136a89ab 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -160,8 +160,8 @@ module FreeRawMagma (A : Set a) where ; ∙-cong = _∙-cong_ } - freeMagma : Magma a a - freeMagma = record { isMagma = isMagma } + magma : Magma a a + magma = record { isMagma = isMagma } ------------------------------------------------------------------------ @@ -169,7 +169,8 @@ module FreeRawMagma (A : Set a) where module FreeMagma (𝓐 : Setoid a ℓa) where - module A = Setoid 𝓐 + private + module A = Setoid 𝓐 open Syntax @@ -184,12 +185,12 @@ module FreeMagma (𝓐 : Setoid a ℓa) where ; ∙-cong = _∙_ } - freeMagma : Magma a (a ⊔ ℓa) - freeMagma = record { isMagma = isMagma } + magma : Magma a (a ⊔ ℓa) + magma = record { isMagma = isMagma } -- re-export some substructure - open Magma freeMagma public using (rawMagma; setoid; Carrier) + open Magma magma public using (rawMagma; setoid; Carrier) varSetoidHomomorphism : SetoidHomomorphism 𝓐 setoid varSetoidHomomorphism = record @@ -203,7 +204,8 @@ module FreeMagma (𝓐 : Setoid a ℓa) where module _ (𝓜 : Magma m ℓm) where - module M = Magma 𝓜 + private + module M = Magma 𝓜 open Syntax ------------------------------------------------------------------------ @@ -211,7 +213,8 @@ module _ (𝓜 : Magma m ℓm) where module Eval (𝓐 : Setoid a ℓa) where - module A = Setoid 𝓐 + private + module A = Setoid 𝓐 ⟦_⟧_ : Syntax A.Carrier → (A.Carrier → M.Carrier) → M.Carrier ⟦ var a ⟧ η = η a @@ -229,103 +232,103 @@ module _ (𝓜 : Magma m ℓm) where module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) (𝓗 : SetoidHomomorphism 𝓐 (Magma.setoid 𝓜)) where - open Magma 𝓜 - renaming (Carrier to UM; _≈_ to _≈ᴹ_; _∙_ to _∙ᴹ_ - ; setoid to setoidᴹ; rawMagma to rawMagmaᴹ - ; isMagma to isMagmaᴹ) + private + module M = Magma 𝓜 + module A = Setoid 𝓐 + module FA = FreeMagma 𝓐 - open ≈-Reasoning setoidᴹ + open ≈-Reasoning M.setoid open Syntax - open Setoid 𝓐 renaming (Carrier to UA; _≈_ to _≈ᴬ_) - - open Eval 𝓜 𝓐 public - - open FreeMagma 𝓐 renaming (setoid to FA; Carrier to UFA) - open SetoidHomomorphism 𝓗 renaming (⟦_⟧ to η; isRelHomomorphism to hom-η) private - ⟦_⟧ᴹ : UFA → UM - ⟦_⟧ᴹ = ⟦_⟧ η + ⟦_⟧ᴹ : FA.Carrier → M.Carrier + ⟦_⟧ᴹ = ⟦_⟧ η where open Eval 𝓜 𝓐 - open Structures _≈ᴹ_ - open IsMagma isMagmaᴹ renaming (∙-cong to congᴹ) + open Structures M._≈_ + open IsMagma M.isMagma renaming (∙-cong to congᴹ) open IsRelHomomorphism hom-η renaming (cong to cong-η) module Existence where - unfold-⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ᴹ ≈ᴹ alg 𝓜 (map η t) - unfold-⟦ var a ⟧ᴹ = begin η a ∎ - unfold-⟦ s ∙ t ⟧ᴹ = congᴹ unfold-⟦ s ⟧ᴹ unfold-⟦ t ⟧ᴹ + unfold-⟦_⟧ : ∀ t → ⟦ t ⟧ᴹ M.≈ alg 𝓜 (map η t) + unfold-⟦ var a ⟧ = begin η a ∎ + unfold-⟦ s ∙ t ⟧ = congᴹ unfold-⟦ s ⟧ unfold-⟦ t ⟧ - cong-⟦_⟧ : ∀ {s t} → s ≈ t → ⟦ s ⟧ᴹ ≈ᴹ ⟦ t ⟧ᴹ - cong-⟦ var r ⟧ = cong-η r - cong-⟦ s ∙ t ⟧ = congᴹ cong-⟦ s ⟧ cong-⟦ t ⟧ + cong-⟦_⟧ : ∀ {s t} → s FA.≈ t → ⟦ s ⟧ᴹ M.≈ ⟦ t ⟧ᴹ + cong-⟦ FA.var r ⟧ = cong-η r + cong-⟦ s FA.∙ t ⟧ = congᴹ cong-⟦ s ⟧ cong-⟦ t ⟧ - isRelHomomorphismᴹ : IsRelHomomorphism _≈_ _≈ᴹ_ ⟦_⟧ᴹ - isRelHomomorphismᴹ = record { cong = cong-⟦_⟧ } + isRelHomomorphism : IsRelHomomorphism FA._≈_ M._≈_ ⟦_⟧ᴹ + isRelHomomorphism = record { cong = cong-⟦_⟧ } - setoidHomomorphismᴹ : SetoidHomomorphism FA setoidᴹ - setoidHomomorphismᴹ = record { ⟦_⟧ = ⟦_⟧ᴹ ; isRelHomomorphism = isRelHomomorphismᴹ } + setoidHomomorphism : SetoidHomomorphism FA.setoid M.setoid + setoidHomomorphism = record + { ⟦_⟧ = ⟦_⟧ᴹ + ; isRelHomomorphism = isRelHomomorphism + } - isMagmaHomomorphismᴹ : IsMagmaHomomorphism rawMagma rawMagmaᴹ ⟦_⟧ᴹ - isMagmaHomomorphismᴹ = record - { isRelHomomorphism = isRelHomomorphismᴹ - ; homo = λ s t → begin ⟦ s ⟧ᴹ ∙ᴹ ⟦ t ⟧ᴹ ∎ - } + isMagmaHomomorphism : IsMagmaHomomorphism FA.rawMagma M.rawMagma ⟦_⟧ᴹ + isMagmaHomomorphism = record + { isRelHomomorphism = isRelHomomorphism + ; homo = λ s t → M.refl + } - magmaHomomorphismᴹ : MagmaHomomorphism freeMagma 𝓜 - magmaHomomorphismᴹ = record { ⟦_⟧ = ⟦_⟧ᴹ - ; isMagmaHomomorphism = isMagmaHomomorphismᴹ } + magmaHomomorphism : MagmaHomomorphism FA.magma 𝓜 + magmaHomomorphism = record + { ⟦_⟧ = ⟦_⟧ᴹ + ; isMagmaHomomorphism = isMagmaHomomorphism + } record η-MagmaHomomorphism : Set (suc (a ⊔ m ⊔ ℓa ⊔ ℓm)) where field - magmaHomomorphism : MagmaHomomorphism freeMagma 𝓜 + magmaHomomorphism : MagmaHomomorphism FA.magma 𝓜 open MagmaHomomorphism magmaHomomorphism public renaming (homo to ⟦⟧-homo) field - ⟦_⟧∘var≈ᴹη : ∀ a → ⟦ var a ⟧ ≈ᴹ η a + ⟦_⟧∘var≈η : ∀ a → ⟦ var a ⟧ M.≈ η a - ⟦⟧ᴹ-η-MagmaHomomorphism : η-MagmaHomomorphism - ⟦⟧ᴹ-η-MagmaHomomorphism = record - { magmaHomomorphism = Existence.magmaHomomorphismᴹ - ; ⟦_⟧∘var≈ᴹη = Existence.unfold-⟦_⟧ᴹ ∘ var + ⟦⟧-η-MagmaHomomorphism : η-MagmaHomomorphism + ⟦⟧-η-MagmaHomomorphism = record + { magmaHomomorphism = Existence.magmaHomomorphism + ; ⟦_⟧∘var≈η = Existence.unfold-⟦_⟧ ∘ var } module Uniqueness (η-magmaHomomorphism : η-MagmaHomomorphism) where open η-MagmaHomomorphism η-magmaHomomorphism - isUnique⟦_⟧ᴹ : ∀ t → ⟦ t ⟧ ≈ᴹ ⟦ t ⟧ᴹ - isUnique⟦ var a ⟧ᴹ = ⟦ a ⟧∘var≈ᴹη - isUnique⟦ s ∙ t ⟧ᴹ = begin + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ M.≈ ⟦ t ⟧ᴹ + isUnique⟦ var a ⟧ = ⟦ a ⟧∘var≈η + isUnique⟦ s ∙ t ⟧ = begin ⟦ s Syntax.∙ t ⟧ ≈⟨ ⟦⟧-homo s t ⟩ - ⟦ s ⟧ ∙ᴹ ⟦ t ⟧ ≈⟨ congᴹ isUnique⟦ s ⟧ᴹ isUnique⟦ t ⟧ᴹ ⟩ - ⟦ s ⟧ᴹ ∙ᴹ ⟦ t ⟧ᴹ ∎ + ⟦ s ⟧ M.∙ ⟦ t ⟧ ≈⟨ congᴹ isUnique⟦ s ⟧ isUnique⟦ t ⟧ ⟩ + ⟦ s ⟧ᴹ M.∙ ⟦ t ⟧ᴹ ∎ module Corollary (𝓗 𝓚 : η-MagmaHomomorphism) where open η-MagmaHomomorphism 𝓗 renaming (⟦_⟧ to ⟦_⟧ᴴ) open η-MagmaHomomorphism 𝓚 renaming (⟦_⟧ to ⟦_⟧ᴷ) - open Uniqueness 𝓗 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴴ) - open Uniqueness 𝓚 renaming (isUnique⟦_⟧ᴹ to isUnique⟦_⟧ᴷ) + open Uniqueness 𝓗 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴴ) + open Uniqueness 𝓚 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴷ) - isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ ≈ᴹ ⟦ t ⟧ᴷ + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ M.≈ ⟦ t ⟧ᴷ isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧ᴹ ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ ------------------------------------------------------------------------ -- Immediate corollary: alg is in fact a MagmaHomomorphism module _ (𝓜 : Magma m ℓm) where - open Magma 𝓜 renaming (setoid to setoidᴹ; _≈_ to _≈ᴹ_; isMagma to isMagmaᴹ) - open FreeMagma setoidᴹ - algMagmaHomomorphism : MagmaHomomorphism freeMagma 𝓜 - algMagmaHomomorphism = Existence.magmaHomomorphismᴹ - where open LeftAdjoint 𝓜 (Identity.setoidHomomorphism setoidᴹ) + private + module M = Magma 𝓜 + + algMagmaHomomorphism : MagmaHomomorphism (FreeMagma.magma M.setoid) 𝓜 + algMagmaHomomorphism = Existence.magmaHomomorphism + where open LeftAdjoint 𝓜 (Identity.setoidHomomorphism M.setoid) ------------------------------------------------------------------------ @@ -333,50 +336,51 @@ module _ (𝓜 : Magma m ℓm) where module FreeMagmaFunctor {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} (𝓗 : SetoidHomomorphism 𝓐 𝓑) where - - open FreeMagma 𝓐 renaming (freeMagma to freeMagmaᴬ) - open FreeMagma 𝓑 renaming (freeMagma to freeMagmaᴮ - ; varSetoidHomomorphism to 𝓥ᴮ) - - mapMagmaHomomorphism : MagmaHomomorphism freeMagmaᴬ freeMagmaᴮ - mapMagmaHomomorphism = Existence.magmaHomomorphismᴹ - where open LeftAdjoint freeMagmaᴮ (Compose.setoidHomomorphism 𝓗 𝓥ᴮ) + private + module FA = FreeMagma 𝓐 + module FB = FreeMagma 𝓑 + + mapMagmaHomomorphism : MagmaHomomorphism FA.magma FB.magma + mapMagmaHomomorphism = Existence.magmaHomomorphism + where open LeftAdjoint FB.magma (Compose.setoidHomomorphism 𝓗 FB.varSetoidHomomorphism) ------------------------------------------------------------------------ -- Naturality of alg module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where - open Magma 𝓜 using () renaming (setoid to setoidᴹ) - open Magma 𝓝 using () renaming (_≈_ to _≈ᴺ_; refl to reflᴺ; trans to transᴺ) + + private + module M = Magma 𝓜 + module N = Magma 𝓝 module _ (𝓕 : MagmaHomomorphism 𝓜 𝓝) where open MagmaHomomorphism 𝓕 using (⟦_⟧; isMagmaHomomorphism; setoidHomomorphism) open FreeMagmaFunctor setoidHomomorphism using (mapMagmaHomomorphism) open MagmaHomomorphism mapMagmaHomomorphism renaming (⟦_⟧ to map; isMagmaHomomorphism to map-isMagmaHomomorphism; setoidHomomorphism to mapSetoidHomomorphism) - open FreeMagma setoidᴹ renaming (freeMagma to freeMagmaᴹ) + open FreeMagma M.setoid renaming (magma to magmaᴹ) open LeftAdjoint 𝓝 setoidHomomorphism open MagmaHomomorphism (algMagmaHomomorphism 𝓜) renaming (⟦_⟧ to algᴹ; isMagmaHomomorphism to algᴹ-isMagmaHomomorphism) open MagmaHomomorphism (algMagmaHomomorphism 𝓝) renaming (⟦_⟧ to algᴺ; isMagmaHomomorphism to algᴺ-isMagmaHomomorphism) - naturality : ∀ t → ⟦ algᴹ t ⟧ ≈ᴺ algᴺ (map t) + naturality : ∀ t → ⟦ algᴹ t ⟧ N.≈ algᴺ (map t) naturality = Corollary.isUnique⟦_⟧ 𝓗 𝓚 where - H K : MagmaHomomorphism freeMagmaᴹ 𝓝 + H K : MagmaHomomorphism magmaᴹ 𝓝 H = record { ⟦_⟧ = ⟦_⟧ ∘ algᴹ - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ algᴹ-isMagmaHomomorphism isMagmaHomomorphism } + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans algᴹ-isMagmaHomomorphism isMagmaHomomorphism } K = record { ⟦_⟧ = algᴺ ∘ map - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transᴺ map-isMagmaHomomorphism algᴺ-isMagmaHomomorphism + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans map-isMagmaHomomorphism algᴺ-isMagmaHomomorphism } 𝓗 𝓚 : η-MagmaHomomorphism - 𝓗 = record { magmaHomomorphism = H ; ⟦_⟧∘var≈ᴹη = λ _ → reflᴺ } - 𝓚 = record { magmaHomomorphism = K ; ⟦_⟧∘var≈ᴹη = λ _ → reflᴺ } + 𝓗 = record { magmaHomomorphism = H ; ⟦_⟧∘var≈η = λ _ → N.refl } + 𝓚 = record { magmaHomomorphism = K ; ⟦_⟧∘var≈η = λ _ → N.refl } ------------------------------------------------------------------------ @@ -384,25 +388,27 @@ module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where module IdentityLaw (𝓐 : Setoid a ℓa) where - open FreeMagma 𝓐 renaming (varSetoidHomomorphism to 𝓥) - open Setoid setoid renaming (_≈_ to _≈FA_; refl to reflFA) + private + module A = Setoid 𝓐 + module FA = FreeMagma 𝓐 + module UFA = Setoid FA.setoid - Id : MagmaHomomorphism freeMagma freeMagma + Id : MagmaHomomorphism FA.magma FA.magma Id = record { ⟦_⟧ = id - ; isMagmaHomomorphism = Identity.isMagmaHomomorphism rawMagma reflFA + ; isMagmaHomomorphism = Identity.isMagmaHomomorphism FA.rawMagma UFA.refl } open FreeMagmaFunctor (Identity.setoidHomomorphism 𝓐) open MagmaHomomorphism mapMagmaHomomorphism renaming (⟦_⟧ to map-Id) - map-id : ∀ t → map-Id t ≈FA t + map-id : ∀ t → map-Id t UFA.≈ t map-id = Corollary.isUnique⟦_⟧ 𝓘ᴬ 𝓘 where - open LeftAdjoint freeMagma 𝓥 + open LeftAdjoint FA.magma FA.varSetoidHomomorphism 𝓘ᴬ 𝓘 : η-MagmaHomomorphism - 𝓘ᴬ = record { magmaHomomorphism = mapMagmaHomomorphism ; ⟦_⟧∘var≈ᴹη = λ _ → reflFA } - 𝓘 = record { magmaHomomorphism = Id ; ⟦_⟧∘var≈ᴹη = λ _ → reflFA } + 𝓘ᴬ = record { magmaHomomorphism = mapMagmaHomomorphism ; ⟦_⟧∘var≈η = λ _ → UFA.refl } + 𝓘 = record { magmaHomomorphism = Id ; ⟦_⟧∘var≈η = λ _ → UFA.refl } module CompositionLaw {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} {𝓒 : Setoid c ℓc} @@ -411,33 +417,32 @@ module CompositionLaw 𝓕 : SetoidHomomorphism 𝓐 𝓒 𝓕 = Compose.setoidHomomorphism 𝓗 𝓚 - open FreeMagma 𝓐 renaming (freeMagma to freeMagmaA) - open FreeMagma 𝓑 renaming (freeMagma to freeMagmaB) - open FreeMagma 𝓒 renaming (freeMagma to freeMagmaC - ; setoid to setoidFC - ; varSetoidHomomorphism to 𝓥) - open Setoid setoidFC renaming (_≈_ to _≈FC_; refl to reflFC; trans to transFC) - 𝓥∘𝓕 = Compose.setoidHomomorphism 𝓕 𝓥 + private + module FA = FreeMagma 𝓐 + module FB = FreeMagma 𝓑 + module FC = FreeMagma 𝓒 + module UFC = Setoid FC.setoid + open FreeMagmaFunctor 𝓕 renaming (mapMagmaHomomorphism to MapAC) open FreeMagmaFunctor 𝓗 renaming (mapMagmaHomomorphism to MapAB) open FreeMagmaFunctor 𝓚 renaming (mapMagmaHomomorphism to MapBC) open MagmaHomomorphism MapAC renaming (⟦_⟧ to mapAC) - open MagmaHomomorphism MapAB renaming (⟦_⟧ to mapAB; isMagmaHomomorphism to isMagmaAB) - open MagmaHomomorphism MapBC renaming (⟦_⟧ to mapBC; isMagmaHomomorphism to isMagmaBC) + open MagmaHomomorphism MapAB renaming (⟦_⟧ to mapAB; isMagmaHomomorphism to isHomAB) + open MagmaHomomorphism MapBC renaming (⟦_⟧ to mapBC; isMagmaHomomorphism to isHomBC) - MapBC∘MapAB : MagmaHomomorphism freeMagmaA freeMagmaC + MapBC∘MapAB : MagmaHomomorphism FA.magma FC.magma MapBC∘MapAB = record { ⟦_⟧ = mapBC ∘ mapAB - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism transFC isMagmaAB isMagmaBC + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism UFC.trans isHomAB isHomBC } - map-∘ : ∀ t → mapAC t ≈FC mapBC (mapAB t) + map-∘ : ∀ t → mapAC t UFC.≈ mapBC (mapAB t) map-∘ = Corollary.isUnique⟦_⟧ 𝓐𝓒 𝓑𝓒∘𝓐𝓑 where - open LeftAdjoint freeMagmaC 𝓥∘𝓕 + open LeftAdjoint FC.magma (Compose.setoidHomomorphism 𝓕 FC.varSetoidHomomorphism) 𝓐𝓒 𝓑𝓒∘𝓐𝓑 : η-MagmaHomomorphism - 𝓐𝓒 = record { magmaHomomorphism = MapAC ; ⟦_⟧∘var≈ᴹη = λ _ → reflFC } - 𝓑𝓒∘𝓐𝓑 = record { magmaHomomorphism = MapBC∘MapAB ; ⟦_⟧∘var≈ᴹη = λ _ → reflFC } + 𝓐𝓒 = record { magmaHomomorphism = MapAC ; ⟦_⟧∘var≈η = λ _ → UFC.refl } + 𝓑𝓒∘𝓐𝓑 = record { magmaHomomorphism = MapBC∘MapAB ; ⟦_⟧∘var≈η = λ _ → UFC.refl } ------------------------------------------------------------------------ diff --git a/src/Algebra/Morphism/Bundles.agda b/src/Algebra/Morphism/Bundles.agda index 527443bbde..7af15f3b80 100644 --- a/src/Algebra/Morphism/Bundles.agda +++ b/src/Algebra/Morphism/Bundles.agda @@ -24,8 +24,9 @@ private ------------------------------------------------------------------------ record MagmaHomomorphism (𝓐 : Magma a ℓa) (𝓑 : Magma b ℓb) : Set (a ⊔ b ⊔ ℓa ⊔ ℓb) where - module A = Magma 𝓐 - module B = Magma 𝓑 + private + module A = Magma 𝓐 + module B = Magma 𝓑 field ⟦_⟧ : A.Carrier → B.Carrier From 20579b5093b8dc233b00ba7a38beaf96979ba125 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 17:02:35 +0100 Subject: [PATCH 15/27] further refactoring to use `private` modules --- src/Algebra/Free/Magma.agda | 46 ++++++++++++++++++------------------- 1 file changed, 23 insertions(+), 23 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 4c136a89ab..febfc1d556 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -245,8 +245,8 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) private - ⟦_⟧ᴹ : FA.Carrier → M.Carrier - ⟦_⟧ᴹ = ⟦_⟧ η where open Eval 𝓜 𝓐 + ⟦_⟧η : FA.Carrier → M.Carrier + ⟦_⟧η = ⟦_⟧ η where open Eval 𝓜 𝓐 open Structures M._≈_ open IsMagma M.isMagma renaming (∙-cong to congᴹ) @@ -254,24 +254,24 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) module Existence where - unfold-⟦_⟧ : ∀ t → ⟦ t ⟧ᴹ M.≈ alg 𝓜 (map η t) + unfold-⟦_⟧ : ∀ t → ⟦ t ⟧η M.≈ alg 𝓜 (map η t) unfold-⟦ var a ⟧ = begin η a ∎ unfold-⟦ s ∙ t ⟧ = congᴹ unfold-⟦ s ⟧ unfold-⟦ t ⟧ - cong-⟦_⟧ : ∀ {s t} → s FA.≈ t → ⟦ s ⟧ᴹ M.≈ ⟦ t ⟧ᴹ + cong-⟦_⟧ : ∀ {s t} → s FA.≈ t → ⟦ s ⟧η M.≈ ⟦ t ⟧η cong-⟦ FA.var r ⟧ = cong-η r cong-⟦ s FA.∙ t ⟧ = congᴹ cong-⟦ s ⟧ cong-⟦ t ⟧ - isRelHomomorphism : IsRelHomomorphism FA._≈_ M._≈_ ⟦_⟧ᴹ + isRelHomomorphism : IsRelHomomorphism FA._≈_ M._≈_ ⟦_⟧η isRelHomomorphism = record { cong = cong-⟦_⟧ } setoidHomomorphism : SetoidHomomorphism FA.setoid M.setoid setoidHomomorphism = record - { ⟦_⟧ = ⟦_⟧ᴹ + { ⟦_⟧ = ⟦_⟧η ; isRelHomomorphism = isRelHomomorphism } - isMagmaHomomorphism : IsMagmaHomomorphism FA.rawMagma M.rawMagma ⟦_⟧ᴹ + isMagmaHomomorphism : IsMagmaHomomorphism FA.rawMagma M.rawMagma ⟦_⟧η isMagmaHomomorphism = record { isRelHomomorphism = isRelHomomorphism ; homo = λ s t → M.refl @@ -279,7 +279,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) magmaHomomorphism : MagmaHomomorphism FA.magma 𝓜 magmaHomomorphism = record - { ⟦_⟧ = ⟦_⟧ᴹ + { ⟦_⟧ = ⟦_⟧η ; isMagmaHomomorphism = isMagmaHomomorphism } @@ -293,30 +293,30 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) ⟦⟧-η-MagmaHomomorphism : η-MagmaHomomorphism ⟦⟧-η-MagmaHomomorphism = record - { magmaHomomorphism = Existence.magmaHomomorphism - ; ⟦_⟧∘var≈η = Existence.unfold-⟦_⟧ ∘ var - } + { magmaHomomorphism = Existence.magmaHomomorphism + ; ⟦_⟧∘var≈η = Existence.unfold-⟦_⟧ ∘ var + } module Uniqueness (η-magmaHomomorphism : η-MagmaHomomorphism) where open η-MagmaHomomorphism η-magmaHomomorphism - isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ M.≈ ⟦ t ⟧ᴹ + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ M.≈ ⟦ t ⟧η isUnique⟦ var a ⟧ = ⟦ a ⟧∘var≈η isUnique⟦ s ∙ t ⟧ = begin - ⟦ s Syntax.∙ t ⟧ ≈⟨ ⟦⟧-homo s t ⟩ - ⟦ s ⟧ M.∙ ⟦ t ⟧ ≈⟨ congᴹ isUnique⟦ s ⟧ isUnique⟦ t ⟧ ⟩ - ⟦ s ⟧ᴹ M.∙ ⟦ t ⟧ᴹ ∎ + ⟦ s Syntax.∙ t ⟧ ≈⟨ ⟦⟧-homo s t ⟩ + ⟦ s ⟧ M.∙ ⟦ t ⟧ ≈⟨ congᴹ isUnique⟦ s ⟧ isUnique⟦ t ⟧ ⟩ + ⟦ s ⟧η M.∙ ⟦ t ⟧η ∎ - module Corollary (𝓗 𝓚 : η-MagmaHomomorphism) - where - open η-MagmaHomomorphism 𝓗 renaming (⟦_⟧ to ⟦_⟧ᴴ) - open η-MagmaHomomorphism 𝓚 renaming (⟦_⟧ to ⟦_⟧ᴷ) - open Uniqueness 𝓗 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴴ) - open Uniqueness 𝓚 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴷ) + module Corollary (𝓗 𝓚 : η-MagmaHomomorphism) where + + open η-MagmaHomomorphism 𝓗 using () renaming (⟦_⟧ to ⟦_⟧ᴴ) + open η-MagmaHomomorphism 𝓚 using () renaming (⟦_⟧ to ⟦_⟧ᴷ) + open Uniqueness 𝓗 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴴ) + open Uniqueness 𝓚 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴷ) - isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ M.≈ ⟦ t ⟧ᴷ - isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧ᴹ ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ + isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ M.≈ ⟦ t ⟧ᴷ + isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧η ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ ------------------------------------------------------------------------ -- Immediate corollary: alg is in fact a MagmaHomomorphism From 77962663413186c55b82a344e9d6fb369d966ead Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 17:04:14 +0100 Subject: [PATCH 16/27] tidied up bugs --- src/Algebra/Free/Magma.agda | 17 ++++++++--------- 1 file changed, 8 insertions(+), 9 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index febfc1d556..761834cc46 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -340,8 +340,8 @@ module FreeMagmaFunctor {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} module FA = FreeMagma 𝓐 module FB = FreeMagma 𝓑 - mapMagmaHomomorphism : MagmaHomomorphism FA.magma FB.magma - mapMagmaHomomorphism = Existence.magmaHomomorphism + magmaHomomorphism : MagmaHomomorphism FA.magma FB.magma + magmaHomomorphism = Existence.magmaHomomorphism where open LeftAdjoint FB.magma (Compose.setoidHomomorphism 𝓗 FB.varSetoidHomomorphism) ------------------------------------------------------------------------ @@ -355,8 +355,7 @@ module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where module _ (𝓕 : MagmaHomomorphism 𝓜 𝓝) where open MagmaHomomorphism 𝓕 using (⟦_⟧; isMagmaHomomorphism; setoidHomomorphism) - open FreeMagmaFunctor setoidHomomorphism using (mapMagmaHomomorphism) - open MagmaHomomorphism mapMagmaHomomorphism + open MagmaHomomorphism (FreeMagmaFunctor.magmaHomomorphism setoidHomomorphism) renaming (⟦_⟧ to map; isMagmaHomomorphism to map-isMagmaHomomorphism; setoidHomomorphism to mapSetoidHomomorphism) open FreeMagma M.setoid renaming (magma to magmaᴹ) open LeftAdjoint 𝓝 setoidHomomorphism @@ -400,14 +399,14 @@ module IdentityLaw (𝓐 : Setoid a ℓa) where } open FreeMagmaFunctor (Identity.setoidHomomorphism 𝓐) - open MagmaHomomorphism mapMagmaHomomorphism renaming (⟦_⟧ to map-Id) + open MagmaHomomorphism magmaHomomorphism renaming (⟦_⟧ to map-Id) map-id : ∀ t → map-Id t UFA.≈ t map-id = Corollary.isUnique⟦_⟧ 𝓘ᴬ 𝓘 where open LeftAdjoint FA.magma FA.varSetoidHomomorphism 𝓘ᴬ 𝓘 : η-MagmaHomomorphism - 𝓘ᴬ = record { magmaHomomorphism = mapMagmaHomomorphism ; ⟦_⟧∘var≈η = λ _ → UFA.refl } + 𝓘ᴬ = record { magmaHomomorphism = magmaHomomorphism ; ⟦_⟧∘var≈η = λ _ → UFA.refl } 𝓘 = record { magmaHomomorphism = Id ; ⟦_⟧∘var≈η = λ _ → UFA.refl } module CompositionLaw @@ -423,9 +422,9 @@ module CompositionLaw module FC = FreeMagma 𝓒 module UFC = Setoid FC.setoid - open FreeMagmaFunctor 𝓕 renaming (mapMagmaHomomorphism to MapAC) - open FreeMagmaFunctor 𝓗 renaming (mapMagmaHomomorphism to MapAB) - open FreeMagmaFunctor 𝓚 renaming (mapMagmaHomomorphism to MapBC) + open FreeMagmaFunctor 𝓕 renaming (magmaHomomorphism to MapAC) + open FreeMagmaFunctor 𝓗 renaming (magmaHomomorphism to MapAB) + open FreeMagmaFunctor 𝓚 renaming (magmaHomomorphism to MapBC) open MagmaHomomorphism MapAC renaming (⟦_⟧ to mapAC) open MagmaHomomorphism MapAB renaming (⟦_⟧ to mapAB; isMagmaHomomorphism to isHomAB) open MagmaHomomorphism MapBC renaming (⟦_⟧ to mapBC; isMagmaHomomorphism to isHomBC) From 9c7eef0e5f8ead651ec43970e3bb162472f2d20c Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 13 May 2023 18:15:06 +0100 Subject: [PATCH 17/27] last refactoring with private modules --- src/Algebra/Free/Magma.agda | 120 +++++++++++++++++------------------- 1 file changed, 58 insertions(+), 62 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 761834cc46..7b1de113f1 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -40,7 +40,12 @@ private A : Set a B : Set b C : Set c - + 𝓐 : Setoid a ℓa + 𝓑 : Setoid b ℓb + 𝓒 : Setoid c ℓc + 𝓜 : Magma m ℓm + 𝓝 : Magma n ℓn + ------------------------------------------------------------------------ -- Syntax: 'pre'-free algebra @@ -334,8 +339,7 @@ module _ (𝓜 : Magma m ℓm) where ------------------------------------------------------------------------ -- Action of FreeMagma on Setoid homomorphisms -module FreeMagmaFunctor {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} - (𝓗 : SetoidHomomorphism 𝓐 𝓑) where +module FreeMagmaFunctor (𝓗 : SetoidHomomorphism 𝓐 𝓑) where private module FA = FreeMagma 𝓐 module FB = FreeMagma 𝓑 @@ -347,39 +351,36 @@ module FreeMagmaFunctor {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} ------------------------------------------------------------------------ -- Naturality of alg -module Naturality {𝓜 : Magma m ℓm} {𝓝 : Magma n ℓn} where +module Naturality (𝓕 : MagmaHomomorphism 𝓜 𝓝) where private module M = Magma 𝓜 module N = Magma 𝓝 + module F = MagmaHomomorphism 𝓕 + module FM = FreeMagma M.setoid + module AlgM = MagmaHomomorphism (algMagmaHomomorphism 𝓜) + module AlgN = MagmaHomomorphism (algMagmaHomomorphism 𝓝) + module Map = MagmaHomomorphism (FreeMagmaFunctor.magmaHomomorphism F.setoidHomomorphism) + + H K : MagmaHomomorphism FM.magma 𝓝 + H = record + { ⟦_⟧ = F.⟦_⟧ ∘ AlgM.⟦_⟧ + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans AlgM.isMagmaHomomorphism F.isMagmaHomomorphism + } + + K = record + { ⟦_⟧ = AlgN.⟦_⟧ ∘ Map.⟦_⟧ + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans Map.isMagmaHomomorphism AlgN.isMagmaHomomorphism + } + + naturality : ∀ t → F.⟦ AlgM.⟦ t ⟧ ⟧ N.≈ AlgN.⟦ Map.⟦ t ⟧ ⟧ + naturality = Corollary.isUnique⟦_⟧ 𝓗 𝓚 + where + open LeftAdjoint 𝓝 F.setoidHomomorphism + 𝓗 𝓚 : η-MagmaHomomorphism + 𝓗 = record { magmaHomomorphism = H ; ⟦_⟧∘var≈η = λ _ → N.refl } + 𝓚 = record { magmaHomomorphism = K ; ⟦_⟧∘var≈η = λ _ → N.refl } - module _ (𝓕 : MagmaHomomorphism 𝓜 𝓝) where - open MagmaHomomorphism 𝓕 using (⟦_⟧; isMagmaHomomorphism; setoidHomomorphism) - open MagmaHomomorphism (FreeMagmaFunctor.magmaHomomorphism setoidHomomorphism) - renaming (⟦_⟧ to map; isMagmaHomomorphism to map-isMagmaHomomorphism; setoidHomomorphism to mapSetoidHomomorphism) - open FreeMagma M.setoid renaming (magma to magmaᴹ) - open LeftAdjoint 𝓝 setoidHomomorphism - open MagmaHomomorphism (algMagmaHomomorphism 𝓜) - renaming (⟦_⟧ to algᴹ; isMagmaHomomorphism to algᴹ-isMagmaHomomorphism) - open MagmaHomomorphism (algMagmaHomomorphism 𝓝) - renaming (⟦_⟧ to algᴺ; isMagmaHomomorphism to algᴺ-isMagmaHomomorphism) - - naturality : ∀ t → ⟦ algᴹ t ⟧ N.≈ algᴺ (map t) - naturality = Corollary.isUnique⟦_⟧ 𝓗 𝓚 - where - H K : MagmaHomomorphism magmaᴹ 𝓝 - H = record - { ⟦_⟧ = ⟦_⟧ ∘ algᴹ - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans algᴹ-isMagmaHomomorphism isMagmaHomomorphism } - - K = record - { ⟦_⟧ = algᴺ ∘ map - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans map-isMagmaHomomorphism algᴺ-isMagmaHomomorphism - } - - 𝓗 𝓚 : η-MagmaHomomorphism - 𝓗 = record { magmaHomomorphism = H ; ⟦_⟧∘var≈η = λ _ → N.refl } - 𝓚 = record { magmaHomomorphism = K ; ⟦_⟧∘var≈η = λ _ → N.refl } ------------------------------------------------------------------------ @@ -391,56 +392,51 @@ module IdentityLaw (𝓐 : Setoid a ℓa) where module A = Setoid 𝓐 module FA = FreeMagma 𝓐 module UFA = Setoid FA.setoid + module IA = FreeMagmaFunctor (Identity.setoidHomomorphism 𝓐) + module MapId = MagmaHomomorphism IA.magmaHomomorphism - Id : MagmaHomomorphism FA.magma FA.magma - Id = record - { ⟦_⟧ = id - ; isMagmaHomomorphism = Identity.isMagmaHomomorphism FA.rawMagma UFA.refl - } - - open FreeMagmaFunctor (Identity.setoidHomomorphism 𝓐) - open MagmaHomomorphism magmaHomomorphism renaming (⟦_⟧ to map-Id) + Id : MagmaHomomorphism FA.magma FA.magma + Id = record + { ⟦_⟧ = id + ; isMagmaHomomorphism = Identity.isMagmaHomomorphism FA.rawMagma UFA.refl + } - map-id : ∀ t → map-Id t UFA.≈ t + map-id : ∀ t → MapId.⟦ t ⟧ UFA.≈ t map-id = Corollary.isUnique⟦_⟧ 𝓘ᴬ 𝓘 where open LeftAdjoint FA.magma FA.varSetoidHomomorphism 𝓘ᴬ 𝓘 : η-MagmaHomomorphism - 𝓘ᴬ = record { magmaHomomorphism = magmaHomomorphism ; ⟦_⟧∘var≈η = λ _ → UFA.refl } + 𝓘ᴬ = record { magmaHomomorphism = IA.magmaHomomorphism ; ⟦_⟧∘var≈η = λ _ → UFA.refl } 𝓘 = record { magmaHomomorphism = Id ; ⟦_⟧∘var≈η = λ _ → UFA.refl } -module CompositionLaw - {𝓐 : Setoid a ℓa} {𝓑 : Setoid b ℓb} {𝓒 : Setoid c ℓc} - (𝓗 : SetoidHomomorphism 𝓐 𝓑) (𝓚 : SetoidHomomorphism 𝓑 𝓒) where - - 𝓕 : SetoidHomomorphism 𝓐 𝓒 - 𝓕 = Compose.setoidHomomorphism 𝓗 𝓚 +module CompositionLaw (𝓗 : SetoidHomomorphism 𝓐 𝓑) (𝓚 : SetoidHomomorphism 𝓑 𝓒) where private module FA = FreeMagma 𝓐 module FB = FreeMagma 𝓑 module FC = FreeMagma 𝓒 module UFC = Setoid FC.setoid + Free𝓗 = FreeMagmaFunctor.magmaHomomorphism 𝓗 + Free𝓚 = FreeMagmaFunctor.magmaHomomorphism 𝓚 + module MapAB = MagmaHomomorphism Free𝓗 + module MapBC = MagmaHomomorphism Free𝓚 + 𝓕 : SetoidHomomorphism 𝓐 𝓒 + 𝓕 = Compose.setoidHomomorphism 𝓗 𝓚 + Free𝓕 = FreeMagmaFunctor.magmaHomomorphism 𝓕 + module MapAC = MagmaHomomorphism Free𝓕 + + MapBC∘MapAB : MagmaHomomorphism FA.magma FC.magma + MapBC∘MapAB = record + { ⟦_⟧ = MapBC.⟦_⟧ ∘ MapAB.⟦_⟧ + ; isMagmaHomomorphism = Compose.isMagmaHomomorphism UFC.trans MapAB.isMagmaHomomorphism MapBC.isMagmaHomomorphism + } - open FreeMagmaFunctor 𝓕 renaming (magmaHomomorphism to MapAC) - open FreeMagmaFunctor 𝓗 renaming (magmaHomomorphism to MapAB) - open FreeMagmaFunctor 𝓚 renaming (magmaHomomorphism to MapBC) - open MagmaHomomorphism MapAC renaming (⟦_⟧ to mapAC) - open MagmaHomomorphism MapAB renaming (⟦_⟧ to mapAB; isMagmaHomomorphism to isHomAB) - open MagmaHomomorphism MapBC renaming (⟦_⟧ to mapBC; isMagmaHomomorphism to isHomBC) - - MapBC∘MapAB : MagmaHomomorphism FA.magma FC.magma - MapBC∘MapAB = record - { ⟦_⟧ = mapBC ∘ mapAB - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism UFC.trans isHomAB isHomBC - } - - map-∘ : ∀ t → mapAC t UFC.≈ mapBC (mapAB t) + map-∘ : ∀ t → MapAC.⟦ t ⟧ UFC.≈ MapBC.⟦ MapAB.⟦ t ⟧ ⟧ map-∘ = Corollary.isUnique⟦_⟧ 𝓐𝓒 𝓑𝓒∘𝓐𝓑 where open LeftAdjoint FC.magma (Compose.setoidHomomorphism 𝓕 FC.varSetoidHomomorphism) 𝓐𝓒 𝓑𝓒∘𝓐𝓑 : η-MagmaHomomorphism - 𝓐𝓒 = record { magmaHomomorphism = MapAC ; ⟦_⟧∘var≈η = λ _ → UFC.refl } + 𝓐𝓒 = record { magmaHomomorphism = Free𝓕 ; ⟦_⟧∘var≈η = λ _ → UFC.refl } 𝓑𝓒∘𝓐𝓑 = record { magmaHomomorphism = MapBC∘MapAB ; ⟦_⟧∘var≈η = λ _ → UFC.refl } From 5da019ee7c798cd2611794df0622ebff9d18beef Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 14 May 2023 10:35:58 +0100 Subject: [PATCH 18/27] bundled `Identity` morphism for `Magma` --- src/Algebra/Morphism/Bundles.agda | 2 +- src/Algebra/Morphism/Construct/Identity.agda | 16 ++++++++++++++++ 2 files changed, 17 insertions(+), 1 deletion(-) diff --git a/src/Algebra/Morphism/Bundles.agda b/src/Algebra/Morphism/Bundles.agda index 7af15f3b80..45cbce291b 100644 --- a/src/Algebra/Morphism/Bundles.agda +++ b/src/Algebra/Morphism/Bundles.agda @@ -1,7 +1,7 @@ ------------------------------------------------------------------------ -- The Agda standard library -- --- Bundled definitions of homomorphisms between Magmas +-- Bundled definitions of homomorphisms between algebras ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} diff --git a/src/Algebra/Morphism/Construct/Identity.agda b/src/Algebra/Morphism/Construct/Identity.agda index 5ea2571de0..5fa28ec502 100644 --- a/src/Algebra/Morphism/Construct/Identity.agda +++ b/src/Algebra/Morphism/Construct/Identity.agda @@ -9,6 +9,7 @@ module Algebra.Morphism.Construct.Identity where open import Algebra.Bundles +open import Algebra.Morphism.Bundles open import Algebra.Morphism.Structures using ( module MagmaMorphisms ; module MonoidMorphisms @@ -247,3 +248,18 @@ module _ (L : RawLoop c ℓ) (open RawLoop L) (refl : Reflexive _≈_) where { isLoopMonomorphism = isLoopMonomorphism ; surjective = _, refl } + +------------------------------------------------------------------------ +-- Bundled morphisms between Magmas +------------------------------------------------------------------------ + +module _ (M : Magma c ℓ) where + + open Magma M using (rawMagma; refl) + open MagmaMorphisms rawMagma rawMagma + + magmaHomomorphism : MagmaHomomorphism M M + magmaHomomorphism = record + { ⟦_⟧ = id + ; isMagmaHomomorphism = isMagmaHomomorphism rawMagma refl + } From 16bc754cf53c3381b19d409caba7864eef9415d4 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 14 May 2023 10:41:24 +0100 Subject: [PATCH 19/27] use of bundled `Identity` morphism for `Magma` --- src/Algebra/Free/Magma.agda | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 7b1de113f1..5f60a1b519 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -396,10 +396,7 @@ module IdentityLaw (𝓐 : Setoid a ℓa) where module MapId = MagmaHomomorphism IA.magmaHomomorphism Id : MagmaHomomorphism FA.magma FA.magma - Id = record - { ⟦_⟧ = id - ; isMagmaHomomorphism = Identity.isMagmaHomomorphism FA.rawMagma UFA.refl - } + Id = Identity.magmaHomomorphism FA.magma map-id : ∀ t → MapId.⟦ t ⟧ UFA.≈ t map-id = Corollary.isUnique⟦_⟧ 𝓘ᴬ 𝓘 From 7ee5b7a7013391007ff61b6d3e35046ac449b8cd Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 14 May 2023 11:00:40 +0100 Subject: [PATCH 20/27] bundled `Composition` of morphisms for `Magma` --- src/Algebra/Free/Magma.agda | 13 ++++++---- .../Morphism/Construct/Composition.agda | 24 +++++++++++++++++++ src/Algebra/Morphism/Construct/Identity.agda | 5 +++- 3 files changed, 36 insertions(+), 6 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 5f60a1b519..4cdc0e799a 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -395,16 +395,19 @@ module IdentityLaw (𝓐 : Setoid a ℓa) where module IA = FreeMagmaFunctor (Identity.setoidHomomorphism 𝓐) module MapId = MagmaHomomorphism IA.magmaHomomorphism - Id : MagmaHomomorphism FA.magma FA.magma - Id = Identity.magmaHomomorphism FA.magma - map-id : ∀ t → MapId.⟦ t ⟧ UFA.≈ t map-id = Corollary.isUnique⟦_⟧ 𝓘ᴬ 𝓘 where open LeftAdjoint FA.magma FA.varSetoidHomomorphism 𝓘ᴬ 𝓘 : η-MagmaHomomorphism - 𝓘ᴬ = record { magmaHomomorphism = IA.magmaHomomorphism ; ⟦_⟧∘var≈η = λ _ → UFA.refl } - 𝓘 = record { magmaHomomorphism = Id ; ⟦_⟧∘var≈η = λ _ → UFA.refl } + 𝓘ᴬ = record + { magmaHomomorphism = IA.magmaHomomorphism + ; ⟦_⟧∘var≈η = λ _ → UFA.refl + } + 𝓘 = record + { magmaHomomorphism = Identity.magmaHomomorphism _ + ; ⟦_⟧∘var≈η = λ _ → UFA.refl + } module CompositionLaw (𝓗 : SetoidHomomorphism 𝓐 𝓑) (𝓚 : SetoidHomomorphism 𝓑 𝓒) where diff --git a/src/Algebra/Morphism/Construct/Composition.agda b/src/Algebra/Morphism/Construct/Composition.agda index 991fb93433..d6e51577a9 100644 --- a/src/Algebra/Morphism/Construct/Composition.agda +++ b/src/Algebra/Morphism/Construct/Composition.agda @@ -9,6 +9,7 @@ module Algebra.Morphism.Construct.Composition where open import Algebra.Bundles +open import Algebra.Morphism.Bundles open import Algebra.Morphism.Structures open import Data.Product open import Function.Base using (_∘_) @@ -384,3 +385,26 @@ module _ {L₁ : RawLoop a ℓ₁} { isLoopMonomorphism = isLoopMonomorphism F.isLoopMonomorphism G.isLoopMonomorphism ; surjective = Func.surjective (_≈_ L₁) (_≈_ L₂) (_≈_ L₃) ≈₃-trans G.⟦⟧-cong F.surjective G.surjective } where module F = IsLoopIsomorphism f-iso; module G = IsLoopIsomorphism g-iso + +------------------------------------------------------------------------ +-- Bundled morphisms between algebras +------------------------------------------------------------------------ + +------------------------------------------------------------------------ +-- Magma + +module _ {M₁ : Magma a ℓ₁} + {M₂ : Magma b ℓ₂} + {M₃ : Magma c ℓ₃} + (f : MagmaHomomorphism M₁ M₂) + (g : MagmaHomomorphism M₂ M₃) + where + + private + module F = MagmaHomomorphism f + module G = MagmaHomomorphism g + + magmaHomomorphism : MagmaHomomorphism M₁ M₃ + magmaHomomorphism = record + { ⟦_⟧ = G.⟦_⟧ ∘ F.⟦_⟧ + ; isMagmaHomomorphism = isMagmaHomomorphism (Magma.trans M₃) F.isMagmaHomomorphism G.isMagmaHomomorphism } diff --git a/src/Algebra/Morphism/Construct/Identity.agda b/src/Algebra/Morphism/Construct/Identity.agda index 5fa28ec502..dd1b074d03 100644 --- a/src/Algebra/Morphism/Construct/Identity.agda +++ b/src/Algebra/Morphism/Construct/Identity.agda @@ -250,9 +250,12 @@ module _ (L : RawLoop c ℓ) (open RawLoop L) (refl : Reflexive _≈_) where } ------------------------------------------------------------------------ --- Bundled morphisms between Magmas +-- Bundled morphisms between algebras ------------------------------------------------------------------------ +------------------------------------------------------------------------ +-- Magma + module _ (M : Magma c ℓ) where open Magma M using (rawMagma; refl) From 325feeaf06e3a38dccc7e1a6163dcca846a4ad0a Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 14 May 2023 11:07:50 +0100 Subject: [PATCH 21/27] use of bundled `Composition` of morphisms for `Magma` --- src/Algebra/Free/Magma.agda | 15 +++++++++------ 1 file changed, 9 insertions(+), 6 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 4cdc0e799a..5de2525893 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -426,18 +426,21 @@ module CompositionLaw (𝓗 : SetoidHomomorphism 𝓐 𝓑) (𝓚 : SetoidHomomo module MapAC = MagmaHomomorphism Free𝓕 MapBC∘MapAB : MagmaHomomorphism FA.magma FC.magma - MapBC∘MapAB = record - { ⟦_⟧ = MapBC.⟦_⟧ ∘ MapAB.⟦_⟧ - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism UFC.trans MapAB.isMagmaHomomorphism MapBC.isMagmaHomomorphism - } + MapBC∘MapAB = Compose.magmaHomomorphism Free𝓗 Free𝓚 map-∘ : ∀ t → MapAC.⟦ t ⟧ UFC.≈ MapBC.⟦ MapAB.⟦ t ⟧ ⟧ map-∘ = Corollary.isUnique⟦_⟧ 𝓐𝓒 𝓑𝓒∘𝓐𝓑 where open LeftAdjoint FC.magma (Compose.setoidHomomorphism 𝓕 FC.varSetoidHomomorphism) 𝓐𝓒 𝓑𝓒∘𝓐𝓑 : η-MagmaHomomorphism - 𝓐𝓒 = record { magmaHomomorphism = Free𝓕 ; ⟦_⟧∘var≈η = λ _ → UFC.refl } - 𝓑𝓒∘𝓐𝓑 = record { magmaHomomorphism = MapBC∘MapAB ; ⟦_⟧∘var≈η = λ _ → UFC.refl } + 𝓐𝓒 = record + { magmaHomomorphism = Free𝓕 + ; ⟦_⟧∘var≈η = λ _ → UFC.refl + } + 𝓑𝓒∘𝓐𝓑 = record + { magmaHomomorphism = MapBC∘MapAB + ; ⟦_⟧∘var≈η = λ _ → UFC.refl + } ------------------------------------------------------------------------ From 3e9cfef0629bbbd3fc3ba05591bd6a5fe8b62712 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 14 May 2023 11:20:01 +0100 Subject: [PATCH 22/27] streamlined use of bundled `Composition` of morphisms --- src/Algebra/Free/Magma.agda | 31 ++++++++++++------------------- 1 file changed, 12 insertions(+), 19 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 5de2525893..5faee612a8 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -357,29 +357,25 @@ module Naturality (𝓕 : MagmaHomomorphism 𝓜 𝓝) where module M = Magma 𝓜 module N = Magma 𝓝 module F = MagmaHomomorphism 𝓕 - module FM = FreeMagma M.setoid + Free𝓕 = FreeMagmaFunctor.magmaHomomorphism (F.setoidHomomorphism) module AlgM = MagmaHomomorphism (algMagmaHomomorphism 𝓜) module AlgN = MagmaHomomorphism (algMagmaHomomorphism 𝓝) - module Map = MagmaHomomorphism (FreeMagmaFunctor.magmaHomomorphism F.setoidHomomorphism) - - H K : MagmaHomomorphism FM.magma 𝓝 - H = record - { ⟦_⟧ = F.⟦_⟧ ∘ AlgM.⟦_⟧ - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans AlgM.isMagmaHomomorphism F.isMagmaHomomorphism - } - - K = record - { ⟦_⟧ = AlgN.⟦_⟧ ∘ Map.⟦_⟧ - ; isMagmaHomomorphism = Compose.isMagmaHomomorphism N.trans Map.isMagmaHomomorphism AlgN.isMagmaHomomorphism - } + + module Map = MagmaHomomorphism Free𝓕 naturality : ∀ t → F.⟦ AlgM.⟦ t ⟧ ⟧ N.≈ AlgN.⟦ Map.⟦ t ⟧ ⟧ naturality = Corollary.isUnique⟦_⟧ 𝓗 𝓚 where open LeftAdjoint 𝓝 F.setoidHomomorphism 𝓗 𝓚 : η-MagmaHomomorphism - 𝓗 = record { magmaHomomorphism = H ; ⟦_⟧∘var≈η = λ _ → N.refl } - 𝓚 = record { magmaHomomorphism = K ; ⟦_⟧∘var≈η = λ _ → N.refl } + 𝓗 = record + { magmaHomomorphism = Compose.magmaHomomorphism (algMagmaHomomorphism 𝓜) 𝓕 + ; ⟦_⟧∘var≈η = λ _ → N.refl + } + 𝓚 = record + { magmaHomomorphism = Compose.magmaHomomorphism Free𝓕 (algMagmaHomomorphism 𝓝) + ; ⟦_⟧∘var≈η = λ _ → N.refl + } @@ -425,9 +421,6 @@ module CompositionLaw (𝓗 : SetoidHomomorphism 𝓐 𝓑) (𝓚 : SetoidHomomo Free𝓕 = FreeMagmaFunctor.magmaHomomorphism 𝓕 module MapAC = MagmaHomomorphism Free𝓕 - MapBC∘MapAB : MagmaHomomorphism FA.magma FC.magma - MapBC∘MapAB = Compose.magmaHomomorphism Free𝓗 Free𝓚 - map-∘ : ∀ t → MapAC.⟦ t ⟧ UFC.≈ MapBC.⟦ MapAB.⟦ t ⟧ ⟧ map-∘ = Corollary.isUnique⟦_⟧ 𝓐𝓒 𝓑𝓒∘𝓐𝓑 where @@ -438,7 +431,7 @@ module CompositionLaw (𝓗 : SetoidHomomorphism 𝓐 𝓑) (𝓚 : SetoidHomomo ; ⟦_⟧∘var≈η = λ _ → UFC.refl } 𝓑𝓒∘𝓐𝓑 = record - { magmaHomomorphism = MapBC∘MapAB + { magmaHomomorphism = Compose.magmaHomomorphism Free𝓗 Free𝓚 ; ⟦_⟧∘var≈η = λ _ → UFC.refl } From a1b5d87a8b320a876e3397f9eeb834807a9542e3 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Fri, 2 Jun 2023 15:53:02 +0100 Subject: [PATCH 23/27] `Effect`ful instances; naming --- src/Algebra/Free/Magma.agda | 13 +++++++++---- 1 file changed, 9 insertions(+), 4 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 5faee612a8..8bfd2fca48 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -17,6 +17,7 @@ open import Algebra.Morphism.Bundles using (MagmaHomomorphism) open import Algebra.Morphism.Structures using (IsMagmaHomomorphism) import Algebra.Morphism.Construct.Identity as Identity import Algebra.Morphism.Construct.Composition as Compose +open import Effect.Applicative open import Effect.Functor open import Effect.Monad open import Function.Base using (id; _∘_) @@ -76,8 +77,8 @@ module Syntax where map-∘ g f (var a) = ≡-refl map-∘ g f (s ∙ t) = (map-∘ g f s) ∙-cong (map-∘ g f t) - syntaxRawFunctor : RawFunctor (Syntax {a}) - syntaxRawFunctor = record { _<$>_ = map } + rawFunctor : RawFunctor (Syntax {a}) + rawFunctor = record { _<$>_ = map } -- Monad instance @@ -85,9 +86,13 @@ module Syntax where bind (var x) h = h x bind (s ∙ t) h = bind s h ∙ bind t h - syntaxRawMonad : RawMonad (Syntax {a}) - syntaxRawMonad = mkRawMonad Syntax var bind + rawMonad : RawMonad (Syntax {a}) + rawMonad = mkRawMonad Syntax var bind +-- Applicative instance + + rawApplicative : RawApplicative (Syntax {a}) + rawApplicative = RawMonad.rawApplicative rawMonad ------------------------------------------------------------------------ -- parametrised 'equational' theory over the 'pre'-free algebra From ff9434e53ee1e4dcbc7421ebcdd2f69efa98efe0 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 2 Jan 2024 17:39:54 +0000 Subject: [PATCH 24/27] updated `CHANGELOG` --- CHANGELOG.md | 3527 +------------------------------------------------- 1 file changed, 44 insertions(+), 3483 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 4bb23133dc..8be73afc66 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1,3529 +1,90 @@ -Version 2.0-dev +Version 2.1-dev =============== -The library has been tested using Agda 2.6.3. +The library has been tested using Agda 2.6.4 and 2.6.4.1. Highlights ---------- -* A golden testing library in `Test.Golden`. This allows you to run a set - of tests and make sure their output matches an expected `golden` value. - The test runner has many options: filtering tests by name, dumping the - list of failures to a file, timing the runs, coloured output, etc. - Cf. the comments in `Test.Golden` and the standard library's own tests - in `tests/` for documentation on how to use the library. - -* A new tactic `cong!` available from `Tactic.Cong` which automatically - infers the argument to `cong` for you via anti-unification. - -* Improved the `solve` tactic in `Tactic.RingSolver` to work in a much - wider range of situations. - -* Added `⌊log₂_⌋` and `⌈log₂_⌉` on Natural Numbers. - -* A massive refactoring of the unindexed Functor/Applicative/Monad hierarchy - and the MonadReader / MonadState type classes. These are now usable with - instance arguments as demonstrated in the tests/monad examples. - - Bug-fixes --------- -* The following operators were missing a fixity declaration, which has now - been fixed - - ``` - infixr 5 _∷_ (Codata.Guarded.Stream) - infix 4 _[_] (Codata.Guarded.Stream) - infixr 5 _∷_ (Codata.Guarded.Stream.Relation.Binary.Pointwise) - infix 4 _≈∞_ (Codata.Guarded.Stream.Relation.Binary.Pointwise) - infixr 5 _∷_ (Codata.Musical.Colist) - infix 4 _≈_ (Codata.Musical.Conat) - infixr 5 _∷_ (Codata.Musical.Colist.Bisimilarity) - infixr 5 _∷_ (Codata.Musical.Colist.Relation.Unary.All) - infixr 5 _∷_ (Codata.Sized.Colist) - infixr 5 _∷_ (Codata.Sized.Covec) - infixr 5 _∷_ (Codata.Sized.Cowriter) - infixl 1 _>>=_ (Codata.Sized.Cowriter) - infixr 5 _∷_ (Codata.Sized.Stream) - infixr 5 _∷_ (Codata.Sized.Colist.Bisimilarity) - infix 4 _ℕ≤?_ (Codata.Sized.Conat.Properties) - infixr 5 _∷_ (Codata.Sized.Covec.Bisimilarity) - infixr 5 _∷_ (Codata.Sized.Cowriter.Bisimilarity) - infixr 5 _∷_ (Codata.Sized.Stream.Bisimilarity) - infixr 8 _⇒_ _⊸_ (Data.Container.Core) - infixr -1 _<$>_ _<*>_ (Data.Container.FreeMonad) - infixl 1 _>>=_ (Data.Container.FreeMonad) - infix 5 _▷_ (Data.Container.Indexed) - infix 4 _≈_ (Data.Float.Base) - infixl 7 _⊓′_ (Data.Nat.Base) - infixl 6 _⊔′_ (Data.Nat.Base) - infixr 8 _^_ (Data.Nat.Base) - infix 4 _!≢0 _!*_!≢0 (Data.Nat.Properties) - infix 4 _≃?_ (Data.Rational.Unnormalised.Properties) - infix 4 _≈ₖᵥ_ (Data.Tree.AVL.Map.Membership.Propositional) - infix 4 _<_ (Induction.WellFounded) - infix -1 _$ⁿ_ (Data.Vec.N-ary) - infix 4 _≋_ (Data.Vec.Functional.Relation.Binary.Equality.Setoid) - infix 4 _≟_ (Reflection.AST.Definition) - infix 4 _≡ᵇ_ (Reflection.AST.Literal) - infix 4 _≈?_ _≟_ _≈_ (Reflection.AST.Meta) - infix 4 _≈?_ _≟_ _≈_ (Reflection.AST.Name) - infix 4 _≟-Telescope_ (Reflection.AST.Term) - infix 4 _≟_ (Reflection.AST.Argument.Information) - infix 4 _≟_ (Reflection.AST.Argument.Modality) - infix 4 _≟_ (Reflection.AST.Argument.Quantity) - infix 4 _≟_ (Reflection.AST.Argument.Relevance) - infix 4 _≟_ (Reflection.AST.Argument.Visibility) - infixr 8 _^_ (Function.Endomorphism.Propositional) - infixr 8 _^_ (Function.Endomorphism.Setoid) - infix 4 _≃_ (Function.HalfAdjointEquivalence) - infix 4 _≈_ _≈ᵢ_ _≤_ (Function.Metric.Bundles) - infixl 6 _∙_ (Function.Metric.Bundles) - infix 4 _≈_ (Function.Metric.Nat.Bundles) - infix 3 _←_ _↢_ (Function.Related) - infix 4 _ℕ<_ _ℕ≤infinity _ℕ≤_ (Codata.Sized.Conat) - infix 6 _ℕ+_ _+ℕ_ (Codata.Sized.Conat) - infixl 4 _+ _* (Data.List.Kleene.Base) - infixr 4 _++++_ _+++*_ _*+++_ _*++*_ (Data.List.Kleene.Base) - infix 4 _[_]* _[_]+ (Data.List.Kleene.Base) - infix 4 _≢∈_ (Data.List.Membership.Propositional) - infixr 5 _`∷_ (Data.List.Reflection) - infix 4 _≡?_ (Data.List.Relation.Binary.Equality.DecPropositional) - infixr 5 _++ᵖ_ (Data.List.Relation.Binary.Prefix.Heterogeneous) - infixr 5 _++ˢ_ (Data.List.Relation.Binary.Suffix.Heterogeneous) - infixr 5 _++_ _++[] (Data.List.Relation.Ternary.Appending.Propositional) - infixr 5 _∷=_ (Data.List.Relation.Unary.Any) - infixr 5 _++_ (Data.List.Ternary.Appending) - infixr 2 _×-⇔_ _×-↣_ _×-↞_ _×-↠_ _×-↔_ _×-cong_ (Data.Product.Function.NonDependent.Propositional) - infixr 2 _×-⟶_ (Data.Product.Function.NonDependent.Setoid) - infixr 2 _×-equivalence_ _×-injection_ _×-left-inverse_ (Data.Product.Function.NonDependent.Setoid) - infixr 2 _×-surjection_ _×-inverse_ (Data.Product.Function.NonDependent.Setoid) - infixr 1 _⊎-⇔_ _⊎-↣_ _⊎-↞_ _⊎-↠_ _⊎-↔_ _⊎-cong_ (Data.Sum.Function.Propositional) - infixr 1 _⊎-⟶_ (Data.Sum.Function.Setoid) - infixr 1 _⊎-equivalence_ _⊎-injection_ _⊎-left-inverse_ (Data.Sum.Function.Setoid) - infixr 1 _⊎-surjection_ _⊎-inverse_ (Data.Sum.Function.Setoid) - infix 8 _⁻¹ (Data.Parity.Base) - infixr 5 _`∷_ (Data.Vec.Reflection) - infixr 5 _∷=_ (Data.Vec.Membership.Setoid) - infix 4 _≟H_ _≟N_ (Algebra.Solver.Ring) - infix 4 _≈_ (Relation.Binary.Bundles) - infixl 6 _∩_ (Relation.Binary.Construct.Intersection) - infix 4 _<₋_ (Relation.Binary.Construct.Add.Infimum.Strict) - infix 4 _≈∙_ (Relation.Binary.Construct.Add.Point.Equality) - infix 4 _≤⁺_ _≤⊤⁺ (Relation.Binary.Construct.Add.Supremum.NonStrict) - infixr 5 _∷ʳ_ (Relation.Binary.Construct.Closure.Transitive) - infixr 5 _∷_ (Codata.Guarded.Stream.Relation.Unary.All) - infixr 5 _∷_ (Foreign.Haskell.List.NonEmpty) - infix 4 _≈_ (Function.Metric.Rational.Bundles) - infixl 6 _ℕ+_ (Level.Literals) - infixr 6 _∪_ (Relation.Binary.Construct.Union) - infixl 6 _+ℤ_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) - infix 4 _≉_ _≈ᵢ_ _≤ᵢ_ (Relation.Binary.Indexed.Homogeneous.Bundles) - infixr 5 _∷ᴹ_ _∷⁻¹ᴹ_ (Text.Regex.Search) - ``` - -* In `System.Exit`, the `ExitFailure` constructor is now carrying an integer - rather than a natural. The previous binding was incorrectly assuming that - all exit codes where non-negative. - -* In `/-monoˡ-≤` in `Data.Nat.DivMod` the parameter `o` was implicit but not inferrable. - It has been changed to be explicit. - -* In `+-distrib-/-∣ʳ` in `Data.Nat.DivMod` the parameter `m` was implicit but not inferrable, - while `n` is explicit but inferrable. They have been changed. - -* In `Function.Definitions` the definitions of `Surjection`, `Inverseˡ`, - `Inverseʳ` were not being re-exported correctly and therefore had an unsolved - meta-variable whenever this module was explicitly parameterised. This has - been fixed. - -* Add module `Algebra.Module` that re-exports the contents of - `Algebra.Module.(Definitions/Structures/Bundles)` - -* Various module-like bundles in `Algebra.Module.Bundles` were missing a fixity - declaration for `_*ᵣ_`. This has been fixed. - -* In `Algebra.Definitions.RawSemiring` the record `prime` add `p∤1 : p ∤ 1#` to the field. - -* In `Data.List.Relation.Ternary.Appending.Setoid` we re-export specialised versions of - the constructors using the setoid's carrier set. Prior to that, the constructor were - re-exported in their full generality which would lead to unsolved meta variables at - their use sites. - -* In `Data.Container.FreeMonad`, we give a direct definition of `_⋆_` as an inductive - type rather than encoding it as an instance of `μ`. This ensures Agda notices that - `C ⋆ X` is strictly positive in `X` which in turn allows us to use the free monad - when defining auxiliary (co)inductive types (cf. the `Tap` example in - `README.Data.Container.FreeMonad`). - -* In `Data.Maybe.Base` the fixity declaration of `_<∣>_` was missing. This has been fixed. - Non-backwards compatible changes -------------------------------- -#### Removed deprecated names - -* All modules and names that were deprecated in v1.2 and before have been removed. - -#### Moved `Codata` modules to `Codata.Sized` - -* Due to the change in Agda 2.6.2 where sized types are no longer compatible - with the `--safe` flag, it has become clear that a third variant of codata - will be needed using coinductive records. Therefore all existing modules in - `Codata` which used sized types have been moved inside a new folder named - `Sized`, e.g. `Codata.Stream` has become `Codata.Sized.Stream`. - -#### Moved `Category` modules to `Effect` - -* As observed by Wen Kokke in Issue #1636, it no longer really makes sense - to group the modules which correspond to the variety of concepts of - (effectful) type constructor arising in functional programming (esp. in haskell) - such as `Monad`, `Applicative`, `Functor`, etc, under `Category.*`, as this - obstructs the importing of the `agda-categories` development into the Standard Library, - and moreover needlessly restricts the applicability of categorical concepts to this - (highly specific) mode of use. Correspondingly, client modules grouped under `*.Categorical.*` which exploit such structure for effectful programming have been renamed `*.Effectful`, with the originals being deprecated. - -### Improvements to pretty printing and regexes - -* In `Text.Pretty`, `Doc` is now a record rather than a type alias. This - helps Agda reconstruct the `width` parameter when the module is opened - without it being applied. In particular this allows users to write - width-generic pretty printers and only pick a width when calling the - renderer by using this import style: - ``` - open import Text.Pretty using (Doc; render) - -- ^-- no width parameter for Doc & render - open module Pretty {w} = Text.Pretty w hiding (Doc; render) - -- ^-- implicit width parameter for the combinators - - pretty : Doc w - pretty = ? -- you can use the combinators here and there won't be any - -- issues related to the fact that `w` cannot be reconstructed - -- anymore - - main = do - -- you can now use the same pretty with different widths: - putStrLn $ render 40 pretty - putStrLn $ render 80 pretty - ``` - -* In `Text.Regex.Search` the `searchND` function finding infix matches has - been tweaked to make sure the returned solution is a local maximum in terms - of length. It used to be that we would make the match start as late as - possible and finish as early as possible. It's now the reverse. - - So `[a-zA-Z]+.agdai?` run on "the path _build/Main.agdai corresponds to" - will return "Main.agdai" when it used to be happy to just return "n.agda". - -### Refactoring of algebraic lattice hierarchy - -* In order to improve modularity and consistency with `Relation.Binary.Lattice`, - the structures & bundles for `Semilattice`, `Lattice`, `DistributiveLattice` - & `BooleanAlgebra` have been moved out of the `Algebra` modules and into their - own hierarchy in `Algebra.Lattice`. - -* All submodules, (e.g. `Algebra.Properties.Semilattice` or `Algebra.Morphism.Lattice`) - have been moved to the corresponding place under `Algebra.Lattice` (e.g. - `Algebra.Lattice.Properties.Semilattice` or `Algebra.Lattice.Morphism.Lattice`). See - the `Deprecated modules` section below for full details. - -* Changed definition of `IsDistributiveLattice` and `IsBooleanAlgebra` so that they are - no longer right-biased which hinders compositionality. More concretely, `IsDistributiveLattice` - now has fields: - ```agda - ∨-distrib-∧ : ∨ DistributesOver ∧ - ∧-distrib-∨ : ∧ DistributesOver ∨ - ``` - instead of - ```agda - ∨-distribʳ-∧ : ∨ DistributesOverʳ ∧ - ``` - and `IsBooleanAlgebra` now has fields: - ``` - ∨-complement : Inverse ⊤ ¬ ∨ - ∧-complement : Inverse ⊥ ¬ ∧ - ``` - instead of: - ```agda - ∨-complementʳ : RightInverse ⊤ ¬ ∨ - ∧-complementʳ : RightInverse ⊥ ¬ ∧ - ``` - -* To allow construction of these structures via their old form, smart constructors - have been added to a new module `Algebra.Lattice.Structures.Biased`, which are by - re-exported automatically by `Algebra.Lattice`. For example, if before you wrote: - ```agda - ∧-∨-isDistributiveLattice = record - { isLattice = ∧-∨-isLattice - ; ∨-distribʳ-∧ = ∨-distribʳ-∧ - } - ``` - you can use the smart constructor `isDistributiveLatticeʳʲᵐ` to write: - ```agda - ∧-∨-isDistributiveLattice = isDistributiveLatticeʳʲᵐ (record - { isLattice = ∧-∨-isLattice - ; ∨-distribʳ-∧ = ∨-distribʳ-∧ - }) - ``` - without having to prove full distributivity. - -* Added new `IsBoundedSemilattice`/`BoundedSemilattice` records. - -* Added new aliases `Is(Meet/Join)(Bounded)Semilattice` for `Is(Bounded)Semilattice` - which can be used to indicate meet/join-ness of the original structures. - -#### Removal of the old function hierarchy - -* The switch to the new function hierarchy is complete and the following modules - have been completely switched over to use the new definitions: - ``` - Data.Sum.Function.Setoid - Data.Sum.Function.Propositional - ``` - -* Additionally the following proofs now use the new definitions instead of the old ones: - * `Algebra.Lattice.Properties.BooleanAlgebra` - * `Algebra.Properties.CommutativeMonoid.Sum` - * `Algebra.Properties.Lattice` - * `replace-equality` - * `Data.Fin.Properties` - * `∀-cons-⇔` - * `⊎⇔∃` - * `Data.Fin.Subset.Properties` - * `out⊆-⇔` - * `in⊆in-⇔` - * `out⊂in-⇔` - * `out⊂out-⇔` - * `in⊂in-⇔` - * `x∈⁅y⁆⇔x≡y` - * `∩⇔×` - * `∪⇔⊎` - * `∃-Subset-[]-⇔` - * `∃-Subset-∷-⇔` - * `Data.List.Countdown` - * `empty` - * `Data.List.Fresh.Relation.Unary.Any` - * `⊎⇔Any` - * `Data.List.NonEmpty` - * `Data.List.Relation.Binary.Lex` - * `[]<[]-⇔` - * `∷<∷-⇔` - * `Data.List.Relation.Binary.Sublist.Heterogeneous.Properties` - * `∷⁻¹` - * `∷ʳ⁻¹` - * `Sublist-[x]-bijection` - * `Data.List.Relation.Binary.Sublist.Setoid.Properties` - * `∷⁻¹` - * `∷ʳ⁻¹` - * `[x]⊆xs⤖x∈xs` - * `Data.List.Relation.Unary.Grouped.Properties` - * `Data.Maybe.Relation.Binary.Connected` - * `just-equivalence` - * `Data.Maybe.Relation.Binary.Pointwise` - * `just-equivalence` - * `Data.Maybe.Relation.Unary.All` - * `just-equivalence` - * `Data.Maybe.Relation.Unary.Any` - * `just-equivalence` - * `Data.Nat.Divisibility` - * `m%n≡0⇔n∣m` - * `Data.Nat.Properties` - * `eq?` - * `Data.Vec.N-ary` - * `uncurry-∀ⁿ` - * `uncurry-∃ⁿ` - * `Data.Vec.Relation.Binary.Lex.Core` - * `P⇔[]<[]` - * `∷<∷-⇔` - * `Data.Vec.Relation.Binary.Pointwise.Extensional` - * `equivalent` - * `Pointwise-≡↔≡` - * `Data.Vec.Relation.Binary.Pointwise.Inductive` - * `Pointwise-≡↔≡` - * `Effect.Monad.Partiality` - * `correct` - * `Relation.Binary.Construct.Closure.Reflexive.Properties` - * `⊎⇔Refl` - * `Relation.Binary.Construct.Closure.Transitive` - * `equivalent` - * `Relation.Binary.Reflection` - * `solve₁` - * `Relation.Nullary.Decidable` - * `map` - - -#### Changes to the new function hierarchy - -* The names of the fields in `Function.Bundles` have been - changed from `f`, `g`, `cong₁` and `cong₂` to `to`, `from`, `to-cong`, `from-cong`. - -* The module `Function.Definitions` no longer has two equalities as module arguments, as - they did not interact as intended with the re-exports from `Function.Definitions.(Core1/Core2)`. - The latter have been removed and their definitions folded into `Function.Definitions`. - -* In `Function.Definitions` the types of `Surjective`, `Injective` and `Surjective` - have been changed from: - ``` - Surjective f = ∀ y → ∃ λ x → f x ≈₂ y - Inverseˡ f g = ∀ y → f (g y) ≈₂ y - Inverseʳ f g = ∀ x → g (f x) ≈₁ x - ``` - to: - ``` - Surjective f = ∀ y → ∃ λ x → ∀ {z} → z ≈₁ x → f z ≈₂ y - Inverseˡ f g = ∀ {x y} → y ≈₁ g x → f y ≈₂ x - Inverseʳ f g = ∀ {x y} → y ≈₂ f x → g y ≈₁ x - ``` - This is for several reasons: i) the new definitions compose much more easily, ii) Agda - can better infer the equalities used. - - To ease backwards compatibility: - - the old definitions have been moved to the new names `StrictlySurjective`, - `StrictlyInverseˡ` and `StrictlyInverseʳ`. - - The records in `Function.Structures` and `Function.Bundles` export proofs - of these under the names `strictlySurjective`, `strictlyInverseˡ` and - `strictlyInverseʳ`, - - Conversion functions have been added in both directions to - `Function.Consequences(.Propositional)`. - -#### Proofs of non-zeroness/positivity/negativity as instance arguments - -* Many numeric operations in the library require their arguments to be non-zero, - and various proofs require their arguments to be non-zero/positive/negative etc. - As discussed on the [mailing list](https://lists.chalmers.se/pipermail/agda/2021/012693.html), - the previous way of constructing and passing round these proofs was extremely clunky and lead - to messy and difficult to read code. We have therefore changed every occurrence where we need - a proof of non-zeroness/positivity/etc. to take it as an irrelevant - [instance argument](https://agda.readthedocs.io/en/latest/language/instance-arguments.html). - See the mailing list for a fuller explanation of the motivation and implementation. - -* For example, whereas the type signature of division used to be: - ``` - _/_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} → ℕ - ``` - it is now: - ``` - _/_ : (dividend divisor : ℕ) .{{_ : NonZero divisor}} → ℕ - ``` - which means that as long as an instance of `NonZero n` is in scope then you can write - `m / n` without having to explicitly provide a proof, as instance search will fill it in - for you. The full list of such operations changed is as follows: - - In `Data.Nat.DivMod`: `_/_`, `_%_`, `_div_`, `_mod_` - - In `Data.Nat.Pseudorandom.LCG`: `Generator` - - In `Data.Integer.DivMod`: `_divℕ_`, `_div_`, `_modℕ_`, `_mod_` - - In `Data.Rational`: `mkℚ+`, `normalize`, `_/_`, `1/_` - - In `Data.Rational.Unnormalised`: `_/_`, `1/_`, `_÷_` - -* At the moment, there are 4 different ways such instance arguments can be provided, - listed in order of convenience and clarity: - 1. *Automatic basic instances* - the standard library provides instances based on the constructors of each - numeric type in `Data.X.Base`. For example, `Data.Nat.Base` constrains an instance of `NonZero (suc n)` for any `n` - and `Data.Integer.Base` contains an instance of `NonNegative (+ n)` for any `n`. Consequently, - if the argument is of the required form, these instances will always be filled in by instance search - automatically, e.g. - ``` - 0/n≡0 : 0 / suc n ≡ 0 - ``` - 2. *Take the instance as an argument* - You can provide the instance argument as a parameter to your function - and Agda's instance search will automatically use it in the correct place without you having to - explicitly pass it, e.g. - ``` - 0/n≡0 : .{{_ : NonZero n}} → 0 / n ≡ 0 - ``` - 3. *Define the instance locally* - You can define an instance argument in scope (e.g. in a `where` clause) - and Agda's instance search will again find it automatically, e.g. - ``` - instance - n≢0 : NonZero n - n≢0 = ... - - 0/n≡0 : 0 / n ≡ 0 - ``` - 4. *Pass the instance argument explicitly* - Finally, if all else fails you can pass the - instance argument explicitly into the function using `{{ }}`, e.g. - ``` - 0/n≡0 : ∀ n (n≢0 : NonZero n) → ((0 / n) {{n≢0}}) ≡ 0 - ``` - Suitable constructors for `NonZero`/`Positive` etc. can be found in `Data.X.Base`. - -* A full list of proofs that have changed to use instance arguments is available at the end of this file. - Notable changes to proofs are now discussed below. - -* Previously one of the hacks used in proofs was to explicitly refer to everything in the correct form, - e.g. if the argument `n` had to be non-zero then you would refer to the argument as `suc n` everywhere - instead of `n`, e.g. - ``` - n/n≡1 : ∀ n → suc n / suc n ≡ 1 - ``` - This made the proofs extremely difficult to use if your term wasn't in the right form. - After being updated to use instance arguments instead, the proof above becomes: - ``` - n/n≡1 : ∀ n {{_ : NonZero n}} → n / n ≡ 1 - ``` - However, note that this means that if you passed in the value `x` to these proofs before, then you - will now have to pass in `suc x`. The proofs for which the arguments have changed form in this way - are highlighted in the list at the bottom of the file. - -* Finally, the definition of `_≢0` has been removed from `Data.Rational.Unnormalised.Base` - and the following proofs about it have been removed from `Data.Rational.Unnormalised.Properties`: - ``` - p≄0⇒∣↥p∣≢0 : ∀ p → p ≠ 0ℚᵘ → ℤ.∣ (↥ p) ∣ ≢0 - ∣↥p∣≢0⇒p≄0 : ∀ p → ℤ.∣ (↥ p) ∣ ≢0 → p ≠ 0ℚᵘ - ``` - -### Change in reduction behaviour of rationals - -* Currently arithmetic expressions involving rationals (both normalised and - unnormalised) undergo disastrous exponential normalisation. For example, - `p + q` would often be normalised by Agda to - `(↥ p ℤ.* ↧ q ℤ.+ ↥ q ℤ.* ↧ p) / (↧ₙ p ℕ.* ↧ₙ q)`. While the normalised form - of `p + q + r + s + t + u + v` would be ~700 lines long. This behaviour - often chokes both type-checking and the display of the expressions in the IDE. - -* To avoid this expansion and make non-trivial reasoning about rationals actually feasible: - 1. the records `ℚᵘ` and `ℚ` have both had the `no-eta-equality` flag enabled - 2. definition of arithmetic operations have trivial pattern matching added to - prevent them reducing, e.g. - ```agda - p + q = (↥ p ℤ.* ↧ q ℤ.+ ↥ q ℤ.* ↧ p) / (↧ₙ p ℕ.* ↧ₙ q) - ``` - has been changed to - ``` - p@record{} + q@record{} = (↥ p ℤ.* ↧ q ℤ.+ ↥ q ℤ.* ↧ p) / (↧ₙ p ℕ.* ↧ₙ q) - ``` - -* As a consequence of this, some proofs that relied on this reduction behaviour - or on eta-equality may no longer go through. There are several ways to fix this: - 1. The principled way is to not rely on this reduction behaviour in the first place. - The `Properties` files for rational numbers have been greatly expanded in `v1.7` - and `v2.0`, and we believe most proofs should be able to be built up from existing - proofs contained within these files. - 2. Alternatively, annotating any rational arguments to a proof with either - `@record{}` or `@(mkℚ _ _ _)` should restore the old reduction behaviour for any - terms involving those parameters. - 3. Finally, if the above approaches are not viable then you may be forced to explicitly - use `cong` combined with a lemma that proves the old reduction behaviour. - -### Change to the definition of `LeftCancellative` and `RightCancellative` - -* The definitions of the types for cancellativity in `Algebra.Definitions` previously - made some of their arguments implicit. This was under the assumption that the operators were - defined by pattern matching on the left argument so that Agda could always infer the - argument on the RHS. - -* Although many of the operators defined in the library follow this convention, this is not - always true and cannot be assumed in user's code. - -* Therefore the definitions have been changed as follows to make all their arguments explicit: - - `LeftCancellative _•_` - - From: `∀ x {y z} → (x • y) ≈ (x • z) → y ≈ z` - - To: `∀ x y z → (x • y) ≈ (x • z) → y ≈ z` - - - `RightCancellative _•_` - - From: `∀ {x} y z → (y • x) ≈ (z • x) → y ≈ z` - - To: `∀ x y z → (y • x) ≈ (z • x) → y ≈ z` - - - `AlmostLeftCancellative e _•_` - - From: `∀ {x} y z → ¬ x ≈ e → (x • y) ≈ (x • z) → y ≈ z` - - To: `∀ x y z → ¬ x ≈ e → (x • y) ≈ (x • z) → y ≈ z` - - - `AlmostRightCancellative e _•_` - - From: `∀ {x} y z → ¬ x ≈ e → (y • x) ≈ (z • x) → y ≈ z` - - To: `∀ x y z → ¬ x ≈ e → (y • x) ≈ (z • x) → y ≈ z` - -* Correspondingly some proofs of the above types will need additional arguments passed explicitly. - Instances can easily be fixed by adding additional underscores, e.g. - - `∙-cancelˡ x` to `∙-cancelˡ x _ _` - - `∙-cancelʳ y z` to `∙-cancelʳ _ y z` - -### Change in the definition of `Prime` - -* The definition of `Prime` in `Data.Nat.Primality` was: - ```agda - Prime 0 = ⊥ - Prime 1 = ⊥ - Prime (suc (suc n)) = (i : Fin n) → 2 + toℕ i ∤ 2 + n - ``` - which was very hard to reason about as not only did it involve conversion - to and from the `Fin` type, it also required that the divisor was of the form - `2 + toℕ i`, which has exactly the same problem as the `suc n` hack described - above used for non-zeroness. - -* To make it easier to use, reason about and read, the definition has been - changed to: - ```agda - Prime 0 = ⊥ - Prime 1 = ⊥ - Prime n = ∀ {d} → 2 ≤ d → d < n → d ∤ n - ``` - -### Renaming of `Reflection` modules - -* Under the `Reflection` module, there were various impending name clashes - between the core AST as exposed by Agda and the annotated AST defined in - the library. - -* While the content of the modules remain the same, the modules themselves - have therefore been renamed as follows: - ``` - Reflection.Annotated ↦ Reflection.AnnotatedAST - Reflection.Annotated.Free ↦ Reflection.AnnotatedAST.Free - - Reflection.Abstraction ↦ Reflection.AST.Abstraction - Reflection.Argument ↦ Reflection.AST.Argument - Reflection.Argument.Information ↦ Reflection.AST.Argument.Information - Reflection.Argument.Quantity ↦ Reflection.AST.Argument.Quantity - Reflection.Argument.Relevance ↦ Reflection.AST.Argument.Relevance - Reflection.Argument.Modality ↦ Reflection.AST.Argument.Modality - Reflection.Argument.Visibility ↦ Reflection.AST.Argument.Visibility - Reflection.DeBruijn ↦ Reflection.AST.DeBruijn - Reflection.Definition ↦ Reflection.AST.Definition - Reflection.Instances ↦ Reflection.AST.Instances - Reflection.Literal ↦ Reflection.AST.Literal - Reflection.Meta ↦ Reflection.AST.Meta - Reflection.Name ↦ Reflection.AST.Name - Reflection.Pattern ↦ Reflection.AST.Pattern - Reflection.Show ↦ Reflection.AST.Show - Reflection.Traversal ↦ Reflection.AST.Traversal - Reflection.Universe ↦ Reflection.AST.Universe - - Reflection.TypeChecking.Monad ↦ Reflection.TCM - Reflection.TypeChecking.Monad.Categorical ↦ Reflection.TCM.Categorical - Reflection.TypeChecking.Monad.Format ↦ Reflection.TCM.Format - Reflection.TypeChecking.Monad.Syntax ↦ Reflection.TCM.Instances - Reflection.TypeChecking.Monad.Instances ↦ Reflection.TCM.Syntax - ``` - -* A new module `Reflection.AST` that re-exports the contents of the - submodules has been added. - -### Implementation of division and modulus for `ℤ` - -* The previous implementations of `_divℕ_`, `_div_`, `_modℕ_`, `_mod_` - in `Data.Integer.DivMod` internally converted to the unary `Fin` datatype - resulting in poor performance. The implementation has been updated to - use the corresponding operations from `Data.Nat.DivMod` which are - efficiently implemented using the Haskell backend. - -### Strict functions - -* The module `Strict` has been deprecated in favour of `Function.Strict` - and the definitions of strict application, `_$!_` and `_$!′_`, have been - moved from `Function.Base` to `Function.Strict`. - -* The contents of `Function.Strict` is now re-exported by `Function`. - -### Changes to ring structures - -* Several ring-like structures now have the multiplicative structure defined by - its laws rather than as a substructure, to avoid repeated proofs that the - underlying relation is an equivalence. These are: - * `IsNearSemiring` - * `IsSemiringWithoutOne` - * `IsSemiringWithoutAnnihilatingZero` - * `IsRing` -* To aid with migration, structures matching the old style ones have been added - to `Algebra.Structures.Biased`, with conversion functions: - * `IsNearSemiring*` and `isNearSemiring*` - * `IsSemiringWithoutOne*` and `isSemiringWithoutOne*` - * `IsSemiringWithoutAnnihilatingZero*` and `isSemiringWithoutAnnihilatingZero*` - * `IsRing*` and `isRing*` - -### Proof-irrelevant empty type - -* The definition of ⊥ has been changed to - ```agda - private - data Empty : Set where - - ⊥ : Set - ⊥ = Irrelevant Empty - ``` - in order to make ⊥ proof irrelevant. Any two proofs of `⊥` or of a negated - statements are now *judgmentally* equal to each other. - -* Consequently we have modified the following definitions: - + In `Relation.Nullary.Decidable.Core`, the type of `dec-no` has changed - ```agda - dec-no : (p? : Dec P) → ¬ P → ∃ λ ¬p′ → p? ≡ no ¬p′ - ↦ dec-no : (p? : Dec P) (¬p : ¬ P) → p? ≡ no ¬p - ``` - + In `Relation.Binary.PropositionalEquality`, the type of `≢-≟-identity` has changed - ```agda - ≢-≟-identity : x ≢ y → ∃ λ ¬eq → x ≟ y ≡ no ¬eq - ↦ ≢-≟-identity : (x≢y : x ≢ y) → x ≟ y ≡ no x≢y - ``` - -### Reorganisation of the `Relation.Nullary` hierarchy - -* It was very difficult to use the `Relation.Nullary` modules, as `Relation.Nullary` - contained the basic definitions of negation, decidability etc., and the operations and - proofs were smeared over `Relation.Nullary.(Negation/Product/Sum/Implication etc.)`. - -* In order to fix this: - - the definition of `Dec` and `recompute` have been moved to `Relation.Nullary.Decidable.Core` - - the definition of `Reflects` has been moved to `Relation.Nullary.Reflects` - - the definition of `¬_` has been moved to `Relation.Nullary.Negation.Core` - -* Backwards compatibility has been maintained, as `Relation.Nullary` still re-exports these publicly. - -* The modules: - ``` - Relation.Nullary.Product - Relation.Nullary.Sum - Relation.Nullary.Implication - ``` - have been deprecated and their contents moved to `Relation.Nullary.(Negation/Reflects/Decidable)` - however all their contents is re-exported by `Relation.Nullary` which is the easiest way to access - it now. - -* In order to facilitate this reorganisation the following breaking moves have occurred: - - `¬?` has been moved from `Relation.Nullary.Negation.Core` to `Relation.Nullary.Decidable.Core` - - `¬-reflects` has been moved from `Relation.Nullary.Negation.Core` to `Relation.Nullary.Reflects`. - - `decidable-stable`, `excluded-middle` and `¬-drop-Dec` have been moved from `Relation.Nullary.Negation` - to `Relation.Nullary.Decidable`. - - `fromDec` and `toDec` have been moved from `Data.Sum.Base` to `Data.Sum`. - -### Refactoring of the unindexed Functor/Applicative/Monad hierarchy - -* The unindexed versions are not defined in terms of the named versions anymore - -* The `RawApplicative` and `RawMonad` type classes have been relaxed so that the underlying - functors do not need their domain and codomain to live at the same Set level. - This is needed for level-increasing functors like `IO : Set l → Set (suc l)`. - -* `RawApplicative` is now `RawFunctor + pure + _<*>_` and `RawMonad` is now - `RawApplicative` + `_>>=_` and so `return` is not used anywhere anymore. - This fixes the conflict with `case_return_of` (#356) - This reorganisation means in particular that the functor/applicative of a monad - are not computed using `_>>=_`. This may break proofs. - -* When `F : Set f → Set f` we moreover have a definable join/μ operator - `join : (M : RawMonad F) → F (F A) → F A`. - -* We now have `RawEmpty` and `RawChoice` respectively packing `empty : M A` and - `(<|>) : M A → M A → M A`. `RawApplicativeZero`, `RawAlternative`, `RawMonadZero`, - `RawMonadPlus` are all defined in terms of these. - -* `MonadT T` now returns a `MonadTd` record that packs both a proof that the - `Monad M` transformed by `T` is a monad and that we can `lift` a computation - `M A` to a transformed computation `T M A`. - -* The monad transformer are not mere aliases anymore, they are record-wrapped - which allows constraints such as `MonadIO (StateT S (ReaderT R IO))` to be - discharged by instance arguments. - -* The mtl-style type classes (`MonadState`, `MonadReader`) do not contain a proof - that the underlying functor is a `Monad` anymore. This ensures we do not have - conflicting `Monad M` instances from a pair of `MonadState S M` & `MonadReader R M` - constraints. - -* `MonadState S M` is now defined in terms of - ```agda - gets : (S → A) → M A - modify : (S → S) → M ⊤ - ``` - with `get` and `put` defined as derived notions. - This is needed because `MonadState S M` does not pack a `Monad M` instance anymore - and so we cannot define `modify f` as `get >>= λ s → put (f s)`. - -* `MonadWriter 𝕎 M` is defined similarly: - ```agda - writer : W × A → M A - listen : M A → M (W × A) - pass : M ((W → W) × A) → M A - ``` - with `tell` defined as a derived notion. - Note that `𝕎` is a `RawMonoid`, not a `Set` and `W` is the carrier of the monoid. - -* New modules: - ``` - Algebra.Construct.Initial - Algebra.Construct.Terminal - Data.List.Effectful.Transformer - Data.List.NonEmpty.Effectful.Transformer - Data.Maybe.Effectful.Transformer - Data.Sum.Effectful.Left.Transformer - Data.Sum.Effectful.Right.Transformer - Data.Vec.Effectful.Transformer - Effect.Empty - Effect.Choice - Effect.Monad.Error.Transformer - Effect.Monad.Identity - Effect.Monad.IO - Effect.Monad.IO.Instances - Effect.Monad.Reader.Indexed - Effect.Monad.Reader.Instances - Effect.Monad.Reader.Transformer - Effect.Monad.Reader.Transformer.Base - Effect.Monad.State.Indexed - Effect.Monad.State.Instances - Effect.Monad.State.Transformer - Effect.Monad.State.Transformer.Base - Effect.Monad.Writer - Effect.Monad.Writer.Indexed - Effect.Monad.Writer.Instances - Effect.Monad.Writer.Transformer - Effect.Monad.Writer.Transformer.Base - IO.Effectful - IO.Instances - ``` - -### Changes to triple reasoning interface - -* The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof that the strict - relation is irreflexive. - -* This allows the following new proof combinator: - ```agda - begin-contradiction : (r : x IsRelatedTo x) → {s : True (IsStrict? r)} → A - ``` - that takes a proof that a value is strictly less than itself and then applies the principle of explosion. - -* Specialised versions of this combinator are available in the following derived modules: - ``` - Data.Nat.Properties - Data.Nat.Binary.Properties - Data.Integer.Properties - Data.Rational.Unnormalised.Properties - Data.Rational.Properties - Data.Vec.Relation.Binary.Lex.Strict - Data.Vec.Relation.Binary.Lex.NonStrict - Relation.Binary.Reasoning.StrictPartialOrder - Relation.Binary.Reasoning.PartialOrder - ``` - -### Other - -* In accordance with changes to the flags in Agda 2.6.3, all modules that previously used - the `--without-K` flag now use the `--cubical-compatible` flag instead. - -* To avoid _large indices_ that are by default no longer allowed in Agda 2.6.4, - universe levels have been increased in the following definitions: - - `Data.Star.Decoration.DecoratedWith` - - `Data.Star.Pointer.Pointer` - - `Reflection.AnnotatedAST.Typeₐ` - - `Reflection.AnnotatedAST.AnnotationFun` - -* The first two arguments of `m≡n⇒m-n≡0` (now `i≡j⇒i-j≡0`) in `Data.Integer.Base` - have been made implicit. - -* The relations `_≤_` `_≥_` `_<_` `_>_` in `Data.Fin.Base` have been - generalised so they can now relate `Fin` terms with different indices. - Should be mostly backwards compatible, but very occasionally when proving - properties about the orderings themselves the second index must be provided - explicitly. - -* The argument `xs` in `xs≮[]` in `Data.{List|Vec}.Relation.Binary.Lex.Strict` - introduced in PRs #1648 and #1672 has now been made implicit. - -* Issue #2075 (Johannes Waldmann): wellfoundedness of the lexicographic ordering - on products, defined in `Data.Product.Relation.Binary.Lex.Strict`, no longer - requires the assumption of symmetry for the first equality relation `_≈₁_`, - leading to a new lemma `Induction.WellFounded.Acc-resp-flip-≈`, and refactoring - of the previous proof `Induction.WellFounded.Acc-resp-≈`. - -* The operation `SymClosure` on relations in - `Relation.Binary.Construct.Closure.Symmetric` has been reimplemented - as a data type `SymClosure _⟶_ a b` that is parameterized by the - input relation `_⟶_` (as well as the elements `a` and `b` of the - domain) so that `_⟶_` can be inferred, which it could not from the - previous implementation using the sum type `a ⟶ b ⊎ b ⟶ a`. - -* In `Algebra.Morphism.Structures`, `IsNearSemiringHomomorphism`, - `IsSemiringHomomorphism`, and `IsRingHomomorphism` have been redesigned to - build up from `IsMonoidHomomorphism`, `IsNearSemiringHomomorphism`, and - `IsSemiringHomomorphism` respectively, adding a single property at each step. - This means that they no longer need to have two separate proofs of - `IsRelHomomorphism`. Similarly, `IsLatticeHomomorphism` is now built as - `IsRelHomomorphism` along with proofs that `_∧_` and `_∨_` are homomorphic. - - Also, `⁻¹-homo` in `IsRingHomomorphism` has been renamed to `-‿homo`. - -* Moved definition of `_>>=_` under `Data.Vec.Base` to its submodule `CartesianBind` - in order to keep another new definition of `_>>=_`, located in `DiagonalBind` - which is also a submodule of `Data.Vec.Base`. - -* The functions `split`, `flatten` and `flatten-split` have been removed from - `Data.List.NonEmpty`. In their place `groupSeqs` and `ungroupSeqs` - have been added to `Data.List.NonEmpty.Base` which morally perform the same - operations but without computing the accompanying proofs. The proofs can be - found in `Data.List.NonEmpty.Properties` under the names `groupSeqs-groups` - and `ungroupSeqs` and `groupSeqs`. - -* In `Data.List.Relation.Unary.Grouped.Properties` the proofs `map⁺` and `map⁻` - have had their preconditions weakened so the equivalences no longer require congruence - proofs. - -* The constructors `+0` and `+[1+_]` from `Data.Integer.Base` are no longer - exported by `Data.Rational.Base`. You will have to open `Data.Integer(.Base)` - directly to use them. - -* The names of the (in)equational reasoning combinators defined in the internal - modules `Data.Rational(.Unnormalised).Properties.≤-Reasoning` have been renamed - (issue #1437) to conform with the defined setoid equality `_≃_` on `Rational`s: - ```agda - step-≈ ↦ step-≃ - step-≃˘ ↦ step-≃˘ - ``` - with corresponding associated syntax: - ```agda - _≈⟨_⟩_ ↦ _≃⟨_⟩_ - _≈˘⟨_⟩_ ↦ _≃˘⟨_⟩_ - ``` - NB. It is not possible to rename or deprecate `syntax` declarations, so Agda will - only issue a "Could not parse the application `begin ...` when scope checking" - warning if the old combinators are used. - -* The types of the proofs `pos⇒1/pos`/`1/pos⇒pos` and `neg⇒1/neg`/`1/neg⇒neg` in - `Data.Rational(.Unnormalised).Properties` have been switched, as the previous - naming scheme didn't correctly generalise to e.g. `pos+pos⇒pos`. For example - the types of `pos⇒1/pos`/`1/pos⇒pos` were: - ``` - pos⇒1/pos : ∀ p .{{_ : NonZero p}} .{{Positive (1/ p)}} → Positive p - 1/pos⇒pos : ∀ p .{{_ : Positive p}} → Positive (1/ p) - ``` - but are now: - ``` - pos⇒1/pos : ∀ p .{{_ : Positive p}} → Positive (1/ p) - 1/pos⇒pos : ∀ p .{{_ : NonZero p}} .{{Positive (1/ p)}} → Positive p - ``` -* `Opₗ` and `Opᵣ` have moved from `Algebra.Core` to `Algebra.Module.Core`. - -* In `Data.Nat.GeneralisedArithmetic`, the `s` and `z` arguments to the following - functions have been made explicit: - * `fold-+` - * `fold-k` - * `fold-*` - * `fold-pull` - -* In `Data.Fin.Properties`: - + the `i` argument to `opposite-suc` has been made explicit; - + `pigeonhole` has been strengthened: wlog, we return a proof that - `i < j` rather than a mere `i ≢ j`. - -* In `Data.Sum.Base` the definitions `fromDec` and `toDec` have been moved to `Data.Sum`. - -* In `Data.Vec.Base`: the definitions `init` and `last` have been changed from the `initLast` - view-derived implementation to direct recursive definitions. - -* In `Codata.Guarded.Stream` the following functions have been modified to have simpler definitions: - * `cycle` - * `interleave⁺` - * `cantor` - Furthermore, the direction of interleaving of `cantor` has changed. Precisely, suppose `pair` is the cantor pairing function, then `lookup (pair i j) (cantor xss)` according to the old definition corresponds to `lookup (pair j i) (cantor xss)` according to the new definition. For a concrete example see the one included at the end of the module. - -* Removed `m/n/o≡m/[n*o]` from `Data.Nat.Divisibility` and added a more general - `m/n/o≡m/[n*o]` to `Data.Nat.DivMod` that doesn't require `n * o ∣ m`. - -* Updated `lookup` functions (and variants) to match the conventions adopted in the - `List` module i.e. `lookup` takes its container first and the index (whose type may - depend on the container value) second. - This affects modules: - ``` - Codata.Guarded.Stream - Codata.Guarded.Stream.Relation.Binary.Pointwise - Codata.Musical.Colist.Base - Codata.Musical.Colist.Relation.Unary.Any.Properties - Codata.Musical.Covec - Codata.Musical.Stream - Codata.Sized.Colist - Codata.Sized.Covec - Codata.Sized.Stream - Data.Vec.Relation.Unary.All - Data.Star.Environment - Data.Star.Pointer - Data.Star.Vec - Data.Trie - Data.Trie.NonEmpty - Data.Tree.AVL - Data.Tree.AVL.Indexed - Data.Tree.AVL.Map - Data.Tree.AVL.NonEmpty - Data.Vec.Recursive - Tactic.RingSolver - Tactic.RingSolver.Core.NatSet - ``` - * Moved & renamed from `Data.Vec.Relation.Unary.All` - to `Data.Vec.Relation.Unary.All.Properties`: - ``` - lookup ↦ lookup⁺ - tabulate ↦ lookup⁻ - ``` - * Renamed in `Data.Vec.Relation.Unary.Linked.Properties` - and `Codata.Guarded.Stream.Relation.Binary.Pointwise`: - ``` - lookup ↦ lookup⁺ - ``` - * Added the following new definitions to `Data.Vec.Relation.Unary.All`: - ``` - lookupAny : All P xs → (i : Any Q xs) → (P ∩ Q) (Any.lookup i) - lookupWith : ∀[ P ⇒ Q ⇒ R ] → All P xs → (i : Any Q xs) → R (Any.lookup i) - lookup : All P xs → (∀ {x} → x ∈ₚ xs → P x) - lookupₛ : P Respects _≈_ → All P xs → (∀ {x} → x ∈ xs → P x) - ``` - * `excluded-middle` in `Relation.Nullary.Decidable.Core` has been renamed to - `¬¬-excluded-middle`. - -Major improvements ------------------- - -### Improvements to ring solver tactic - -* The ring solver tactic has been greatly improved. In particular: - 1. When the solver is used for concrete ring types, e.g. ℤ, the equality can now use - all the ring operations defined natively for that type, rather than having - to use the operations defined in `AlmostCommutativeRing`. For example - previously you could not use `Data.Integer.Base._*_` but instead had to - use `AlmostCommutativeRing._*_`. - 2. The solver now supports use of the subtraction operator `_-_` whenever - it is defined immediately in terms of `_+_` and `-_`. This is the case for - `Data.Integer` and `Data.Rational`. - -### Moved `_%_` and `_/_` operators to `Data.Nat.Base` - -* Previously the division and modulus operators were defined in `Data.Nat.DivMod` - which in turn meant that using them required importing `Data.Nat.Properties` - which is a very heavy dependency. - -* To fix this, these operators have been moved to `Data.Nat.Base`. The properties - for them still live in `Data.Nat.DivMod` (which also publicly re-exports them - to provide backwards compatibility). - -* Beneficiaries of this change include `Data.Rational.Unnormalised.Base` whose - dependencies are now significantly smaller. - -### Moved raw bundles from Data.X.Properties to Data.X.Base - -* As mentioned by MatthewDaggitt in Issue #1755, Raw bundles defined in Data.X.Properties - should be defined in Data.X.Base as they don't require any properties. - * Moved raw bundles From `Data.Nat.Properties` to `Data.Nat.Base` - * Moved raw bundles From `Data.Nat.Binary.Properties` to `Data.Nat.Binary.Base` - * Moved raw bundles From `Data.Rational.Properties` to `Data.Rational.Base` - * Moved raw bundles From `Data.Rational.Unnormalised.Properties` to `Data.Rational.Unnormalised.Base` - -### Moved the definition of RawX from `Algebra.X.Bundles` to `Algebra.X.Bundles.Raw` - -* A new module `Algebra.Bundles.Raw` containing the definitions of the raw bundles - can be imported at much lower cost from `Data.X.Base`. - The following definitions have been moved: - * `RawMagma` - * `RawMonoid` - * `RawGroup` - * `RawNearSemiring` - * `RawSemiring` - * `RawRingWithoutOne` - * `RawRing` - * `RawQuasigroup` - * `RawLoop` - * `RawKleeneAlgebra` -* A new module `Algebra.Lattice.Bundles.Raw` is also introduced. - * `RawLattice` has been moved from `Algebra.Lattice.Bundles` to this new module. - -* In `Relation.Binary.Reasoning.Base.Triple`, added a new parameter `<-irrefl : Irreflexive _≈_ _<_` +Other major improvements +------------------------ Deprecated modules ------------------ -### Moving `Relation.Binary.Construct.(Converse/Flip)` +Deprecated names +---------------- -* The following files have been moved: +* In `Data.Nat.Divisibility.Core`: ```agda - Relation.Binary.Construct.Converse ↦ Relation.Binary.Construct.Flip.EqAndOrd - Relation.Binary.Construct.Flip ↦ Relation.Binary.Construct.Flip.Ord + *-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣ ``` -### Deprecation of old function hierarchy - -* The module `Function.Related` has been deprecated in favour of `Function.Related.Propositional` - whose code uses the new function hierarchy. This also opens up the possibility of a more - general `Function.Related.Setoid` at a later date. Several of the names have been changed - in this process to bring them into line with the camelcase naming convention used - in the rest of the library: - ```agda - reverse-implication ↦ reverseImplication - reverse-injection ↦ reverseInjection - left-inverse ↦ leftInverse - - Symmetric-kind ↦ SymmetricKind - Forward-kind ↦ ForwardKind - Backward-kind ↦ BackwardKind - Equivalence-kind ↦ EquivalenceKind - ``` +New modules +----------- -### Moving `Algebra.Lattice` files +Additions to existing modules +----------------------------- -* As discussed above the following files have been moved: +* In `Data.Fin.Properties`: ```agda - Algebra.Properties.Semilattice ↦ Algebra.Lattice.Properties.Semilattice - Algebra.Properties.Lattice ↦ Algebra.Lattice.Properties.Lattice - Algebra.Properties.DistributiveLattice ↦ Algebra.Lattice.Properties.DistributiveLattice - Algebra.Properties.BooleanAlgebra ↦ Algebra.Lattice.Properties.BooleanAlgebra - Algebra.Properties.BooleanAlgebra.Expression ↦ Algebra.Lattice.Properties.BooleanAlgebra.Expression - Algebra.Morphism.LatticeMonomorphism ↦ Algebra.Lattice.Morphism.LatticeMonomorphism + nonZeroIndex : Fin n → ℕ.NonZero n ``` -### Moving `*.Catgeorical.*` files - -* As discussed above the following files have been moved: +* In `Data.List.Relation.Unary.All.Properties`: ```agda - Codata.Sized.Colist.Categorical ↦ Codata.Sized.Colist.Effectful - Codata.Sized.Covec.Categorical ↦ Codata.Sized.Covec.Effectful - Codata.Sized.Delay.Categorical ↦ Codata.Sized.Delay.Effectful - Codata.Sized.Stream.Categorical ↦ Codata.Sized.Stream.Effectful - Data.List.Categorical ↦ Data.List.Effectful - Data.List.Categorical.Transformer ↦ Data.List.Effectful.Transformer - Data.List.NonEmpty.Categorical ↦ Data.List.NonEmpty.Effectful - Data.List.NonEmpty.Categorical.Transformer ↦ Data.List.NonEmpty.Effectful.Transformer - Data.Maybe.Categorical ↦ Data.Maybe.Effectful - Data.Maybe.Categorical.Transformer ↦ Data.Maybe.Effectful.Transformer - Data.Product.Categorical.Examples ↦ Data.Product.Effectful.Examples - Data.Product.Categorical.Left ↦ Data.Product.Effectful.Left - Data.Product.Categorical.Left.Base ↦ Data.Product.Effectful.Left.Base - Data.Product.Categorical.Right ↦ Data.Product.Effectful.Right - Data.Product.Categorical.Right.Base ↦ Data.Product.Effectful.Right.Base - Data.Sum.Categorical.Examples ↦ Data.Sum.Effectful.Examples - Data.Sum.Categorical.Left ↦ Data.Sum.Effectful.Left - Data.Sum.Categorical.Left.Transformer ↦ Data.Sum.Effectful.Left.Transformer - Data.Sum.Categorical.Right ↦ Data.Sum.Effectful.Right - Data.Sum.Categorical.Right.Transformer ↦ Data.Sum.Effectful.Right.Transformer - Data.These.Categorical.Examples ↦ Data.These.Effectful.Examples - Data.These.Categorical.Left ↦ Data.These.Effectful.Left - Data.These.Categorical.Left.Base ↦ Data.These.Effectful.Left.Base - Data.These.Categorical.Right ↦ Data.These.Effectful.Right - Data.These.Categorical.Right.Base ↦ Data.These.Effectful.Right.Base - Data.Vec.Categorical ↦ Data.Vec.Effectful - Data.Vec.Categorical.Transformer ↦ Data.Vec.Effectful.Transformer - Data.Vec.Recursive.Categorical ↦ Data.Vec.Recursive.Effectful - Function.Identity.Categorical ↦ Function.Identity.Effectful - IO.Categorical ↦ IO.Effectful - Reflection.TCM.Categorical ↦ Reflection.TCM.Effectful + All-catMaybes⁺ : All (Maybe.All P) xs → All P (catMaybes xs) + Any-catMaybes⁺ : All (Maybe.Any P) xs → All P (catMaybes xs) ``` -### Moving `Relation.Binary.Properties.XLattice` files - -* The following files have been moved: +* In `Data.List.Relation.Unary.AllPairs.Properties`: ```agda - Relation.Binary.Properties.BoundedJoinSemilattice.agda ↦ Relation.Binary.Lattice.Properties.BoundedJoinSemilattice.agda - Relation.Binary.Properties.BoundedLattice.agda ↦ Relation.Binary.Lattice.Properties.BoundedLattice.agda - Relation.Binary.Properties.BoundedMeetSemilattice.agda ↦ Relation.Binary.Lattice.Properties.BoundedMeetSemilattice.agda - Relation.Binary.Properties.DistributiveLattice.agda ↦ Relation.Binary.Lattice.Properties.DistributiveLattice.agda - Relation.Binary.Properties.JoinSemilattice.agda ↦ Relation.Binary.Lattice.Properties.JoinSemilattice.agda - Relation.Binary.Properties.Lattice.agda ↦ Relation.Binary.Lattice.Properties.Lattice.agda - Relation.Binary.Properties.MeetSemilattice.agda ↦ Relation.Binary.Lattice.Properties.MeetSemilattice.agda + catMaybes⁺ : AllPairs (Pointwise R) xs → AllPairs R (catMaybes xs) + tabulate⁺-< : (i < j → R (f i) (f j)) → AllPairs R (tabulate f) ``` -### Deprecation of `Data.Nat.Properties.Core` - -* The module `Data.Nat.Properties.Core` has been deprecated, and its one entry moved to `Data.Nat.Properties` - -### Deprecation of `Data.Fin.Substitution.Example` - -* The module `Data.Fin.Substitution.Example` has been deprecated, and moved to `README.Data.Fin.Substitution.UntypedLambda` - -### Deprecation of `Data.Product.Function.Dependent.Setoid.WithK` - -* This module has been deprecated, as none of its contents actually depended on axiom K. The contents has been moved to `Data.Product.Function.Dependent.Setoid`. - -Deprecated names ----------------- - -* In `Algebra.Construct.Zero`: +* In `Data.Maybe.Relation.Binary.Pointwise`: ```agda - rawMagma ↦ Algebra.Construct.Terminal.rawMagma - magma ↦ Algebra.Construct.Terminal.magma - semigroup ↦ Algebra.Construct.Terminal.semigroup - band ↦ Algebra.Construct.Terminal.band + pointwise⊆any : Pointwise R (just x) ⊆ Any (R x) ``` -* In `Codata.Guarded.Stream.Properties`: +* In `Data.Nat.Divisibility`: ```agda - map-identity ↦ map-id - map-fusion ↦ map-∘ - drop-fusion ↦ drop-drop - ``` + quotient≢0 : m ∣ n → .{{NonZero n}} → NonZero quotient -* In `Codata.Sized.Colist.Properties`: - ```agda - map-identity ↦ map-id - map-map-fusion ↦ map-∘ - drop-drop-fusion ↦ drop-drop - ``` + m∣n⇒n≡quotient*m : m ∣ n → n ≡ quotient * m + m∣n⇒n≡m*quotient : m ∣ n → n ≡ m * quotient + quotient-∣ : m ∣ n → quotient ∣ n + quotient>1 : m ∣ n → m < n → 1 < quotient + quotient-< : m ∣ n → .{{NonTrivial m}} → .{{NonZero n}} → quotient < n + n/m≡quotient : m ∣ n → .{{_ : NonZero m}} → n / m ≡ quotient -* In `Codata.Sized.Covec.Properties`: - ```agda - map-identity ↦ map-id - map-map-fusion ↦ map-∘ - ``` + m/n≡0⇒m-commute ↦ map-<∣> - -* In `Data.List.Relation.Binary.Subset.Propositional.Properties`: - ``` - map-with-∈⁺ ↦ mapWith∈⁺ - ``` - -* In `Data.List.Relation.Unary.Any.Properties`: - ``` - map-with-∈⁺ ↦ mapWith∈⁺ - map-with-∈⁻ ↦ mapWith∈⁻ - map-with-∈↔ ↦ mapWith∈↔ - ``` - -* In `Data.List.Relation.Unary.All.Properties`: - ``` - gmap ↦ gmap⁺ - ``` - -* In `Data.Nat.Properties`: - ``` - suc[pred[n]]≡n ↦ suc-pred - ≤-step ↦ m≤n⇒m≤1+n - ≤-stepsˡ ↦ m≤n⇒m≤o+n - ≤-stepsʳ ↦ m≤n⇒m≤n+o - <-step ↦ m-idem : Idempotent _<∣>_ - ``` - -* The module `Algebra` now publicly re-exports the contents of - `Algebra.Structures.Biased`. - -* Added new definitions to `Algebra.Bundles`: - ```agda - record UnitalMagma c ℓ : Set (suc (c ⊔ ℓ)) - record InvertibleMagma c ℓ : Set (suc (c ⊔ ℓ)) - record InvertibleUnitalMagma c ℓ : Set (suc (c ⊔ ℓ)) - record RawQuasiGroup c ℓ : Set (suc (c ⊔ ℓ)) - record Quasigroup c ℓ : Set (suc (c ⊔ ℓ)) - record RawLoop c ℓ : Set (suc (c ⊔ ℓ)) - record Loop c ℓ : Set (suc (c ⊔ ℓ)) - record RingWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) - record IdempotentSemiring c ℓ : Set (suc (c ⊔ ℓ)) - record KleeneAlgebra c ℓ : Set (suc (c ⊔ ℓ)) - record RawRingWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) - record Quasiring c ℓ : Set (suc (c ⊔ ℓ)) where - record Nearring c ℓ : Set (suc (c ⊔ ℓ)) where - record IdempotentMagma c ℓ : Set (suc (c ⊔ ℓ)) - record AlternateMagma c ℓ : Set (suc (c ⊔ ℓ)) - record FlexibleMagma c ℓ : Set (suc (c ⊔ ℓ)) - record MedialMagma c ℓ : Set (suc (c ⊔ ℓ)) - record SemimedialMagma c ℓ : Set (suc (c ⊔ ℓ)) - record LeftBolLoop c ℓ : Set (suc (c ⊔ ℓ)) - record RightBolLoop c ℓ : Set (suc (c ⊔ ℓ)) - record MoufangLoop c ℓ : Set (suc (c ⊔ ℓ)) - record NonAssociativeRing c ℓ : Set (suc (c ⊔ ℓ)) - record MiddleBolLoop c ℓ : Set (suc (c ⊔ ℓ)) - ``` - and the existing record `Lattice` now provides - ```agda - ∨-commutativeSemigroup : CommutativeSemigroup c ℓ - ∧-commutativeSemigroup : CommutativeSemigroup c ℓ - ``` - and their corresponding algebraic subbundles. - -* Added new proofs to `Algebra.Consequences.Base`: - ```agda - reflexive+selfInverse⇒involutive : Reflexive _≈_ → - SelfInverse _≈_ f → - Involutive _≈_ f - ``` - -* Added new proofs to `Algebra.Consequences.Propositional`: - ```agda - comm+assoc⇒middleFour : Commutative _•_ → - Associative _•_ → - _•_ MiddleFourExchange _•_ - identity+middleFour⇒assoc : Identity e _•_ → - _•_ MiddleFourExchange _•_ → - Associative _•_ - identity+middleFour⇒comm : Identity e _+_ → - _•_ MiddleFourExchange _+_ → - Commutative _•_ - ``` - -* Added new proofs to `Algebra.Consequences.Setoid`: - ```agda - comm+assoc⇒middleFour : Congruent₂ _•_ → Commutative _•_ → Associative _•_ → - _•_ MiddleFourExchange _•_ - identity+middleFour⇒assoc : Congruent₂ _•_ → Identity e _•_ → - _•_ MiddleFourExchange _•_ → - Associative _•_ - identity+middleFour⇒comm : Congruent₂ _•_ → Identity e _+_ → - _•_ MiddleFourExchange _+_ → - Commutative _•_ - - involutive⇒surjective : Involutive f → Surjective f - selfInverse⇒involutive : SelfInverse f → Involutive f - selfInverse⇒congruent : SelfInverse f → Congruent f - selfInverse⇒inverseᵇ : SelfInverse f → Inverseᵇ f f - selfInverse⇒surjective : SelfInverse f → Surjective f - selfInverse⇒injective : SelfInverse f → Injective f - selfInverse⇒bijective : SelfInverse f → Bijective f - - comm+idˡ⇒id : Commutative _•_ → LeftIdentity e _•_ → Identity e _•_ - comm+idʳ⇒id : Commutative _•_ → RightIdentity e _•_ → Identity e _•_ - comm+zeˡ⇒ze : Commutative _•_ → LeftZero e _•_ → Zero e _•_ - comm+zeʳ⇒ze : Commutative _•_ → RightZero e _•_ → Zero e _•_ - comm+invˡ⇒inv : Commutative _•_ → LeftInverse e _⁻¹ _•_ → Inverse e _⁻¹ _•_ - comm+invʳ⇒inv : Commutative _•_ → RightInverse e _⁻¹ _•_ → Inverse e _⁻¹ _•_ - comm+distrˡ⇒distr : Commutative _•_ → _•_ DistributesOverˡ _◦_ → _•_ DistributesOver _◦_ - comm+distrʳ⇒distr : Commutative _•_ → _•_ DistributesOverʳ _◦_ → _•_ DistributesOver _◦_ - distrib+absorbs⇒distribˡ : Associative _•_ → Commutative _◦_ → _•_ Absorbs _◦_ → _◦_ Absorbs _•_ → _◦_ DistributesOver _•_ → _•_ DistributesOverˡ _◦_ - ``` - -* Added new functions to `Algebra.Construct.DirectProduct`: - ```agda - rawSemiring : RawSemiring a ℓ₁ → RawSemiring b ℓ₂ → RawSemiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - rawRing : RawRing a ℓ₁ → RawRing b ℓ₂ → RawRing (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - semiringWithoutAnnihilatingZero : SemiringWithoutAnnihilatingZero a ℓ₁ → - SemiringWithoutAnnihilatingZero b ℓ₂ → - SemiringWithoutAnnihilatingZero (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - semiring : Semiring a ℓ₁ → Semiring b ℓ₂ → Semiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - commutativeSemiring : CommutativeSemiring a ℓ₁ → CommutativeSemiring b ℓ₂ → - CommutativeSemiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - ring : Ring a ℓ₁ → Ring b ℓ₂ → Ring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - commutativeRing : CommutativeRing a ℓ₁ → CommutativeRing b ℓ₂ → - CommutativeRing (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - rawQuasigroup : RawQuasigroup a ℓ₁ → RawQuasigroup b ℓ₂ → - RawQuasigroup (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - rawLoop : RawLoop a ℓ₁ → RawLoop b ℓ₂ → RawLoop (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - unitalMagma : UnitalMagma a ℓ₁ → UnitalMagma b ℓ₂ → - UnitalMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - invertibleMagma : InvertibleMagma a ℓ₁ → InvertibleMagma b ℓ₂ → - InvertibleMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - invertibleUnitalMagma : InvertibleUnitalMagma a ℓ₁ → InvertibleUnitalMagma b ℓ₂ → - InvertibleUnitalMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - quasigroup : Quasigroup a ℓ₁ → Quasigroup b ℓ₂ → Quasigroup (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - loop : Loop a ℓ₁ → Loop b ℓ₂ → Loop (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - idempotentSemiring : IdempotentSemiring a ℓ₁ → IdempotentSemiring b ℓ₂ → - IdempotentSemiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - idempotentMagma : IdempotentMagma a ℓ₁ → IdempotentMagma b ℓ₂ → - IdempotentMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - alternativeMagma : AlternativeMagma a ℓ₁ → AlternativeMagma b ℓ₂ → - AlternativeMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - flexibleMagma : FlexibleMagma a ℓ₁ → FlexibleMagma b ℓ₂ → - FlexibleMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - medialMagma : MedialMagma a ℓ₁ → MedialMagma b ℓ₂ → MedialMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - semimedialMagma : SemimedialMagma a ℓ₁ → SemimedialMagma b ℓ₂ → - SemimedialMagma (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - kleeneAlgebra : KleeneAlgebra a ℓ₁ → KleeneAlgebra b ℓ₂ → KleeneAlgebra (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - leftBolLoop : LeftBolLoop a ℓ₁ → LeftBolLoop b ℓ₂ → LeftBolLoop (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - rightBolLoop : RightBolLoop a ℓ₁ → RightBolLoop b ℓ₂ → RightBolLoop (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - middleBolLoop : MiddleBolLoop a ℓ₁ → MiddleBolLoop b ℓ₂ → MiddleBolLoop (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - moufangLoop : MoufangLoop a ℓ₁ → MoufangLoop b ℓ₂ → MoufangLoop (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - rawRingWithoutOne : RawRingWithoutOne a ℓ₁ → RawRingWithoutOne b ℓ₂ → RawRingWithoutOne (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - ringWithoutOne : RingWithoutOne a ℓ₁ → RingWithoutOne b ℓ₂ → RingWithoutOne (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - nonAssociativeRing : NonAssociativeRing a ℓ₁ → NonAssociativeRing b ℓ₂ → NonAssociativeRing (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - quasiring : Quasiring a ℓ₁ → Quasiring b ℓ₂ → Quasiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - nearring : Nearring a ℓ₁ → Nearring b ℓ₂ → Nearring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - ``` - -* Added new definition to `Algebra.Definitions`: - ```agda - _MiddleFourExchange_ : Op₂ A → Op₂ A → Set _ - - SelfInverse : Op₁ A → Set _ - - LeftDividesˡ : Op₂ A → Op₂ A → Set _ - LeftDividesʳ : Op₂ A → Op₂ A → Set _ - RightDividesˡ : Op₂ A → Op₂ A → Set _ - RightDividesʳ : Op₂ A → Op₂ A → Set _ - LeftDivides : Op₂ A → Op₂ A → Set _ - RightDivides : Op₂ A → Op₂ A → Set _ - - LeftInvertible e _∙_ x = ∃[ x⁻¹ ] (x⁻¹ ∙ x) ≈ e - RightInvertible e _∙_ x = ∃[ x⁻¹ ] (x ∙ x⁻¹) ≈ e - Invertible e _∙_ x = ∃[ x⁻¹ ] ((x⁻¹ ∙ x) ≈ e) × ((x ∙ x⁻¹) ≈ e) - StarRightExpansive e _+_ _∙_ _⁻* = ∀ x → (e + (x ∙ (x ⁻*))) ≈ (x ⁻*) - StarLeftExpansive e _+_ _∙_ _⁻* = ∀ x → (e + ((x ⁻*) ∙ x)) ≈ (x ⁻*) - StarExpansive e _+_ _∙_ _* = (StarLeftExpansive e _+_ _∙_ _*) × (StarRightExpansive e _+_ _∙_ _*) - StarLeftDestructive _+_ _∙_ _* = ∀ a b x → (b + (a ∙ x)) ≈ x → ((a *) ∙ b) ≈ x - StarRightDestructive _+_ _∙_ _* = ∀ a b x → (b + (x ∙ a)) ≈ x → (b ∙ (a *)) ≈ x - StarDestructive _+_ _∙_ _* = (StarLeftDestructive _+_ _∙_ _*) × (StarRightDestructive _+_ _∙_ _*) - LeftAlternative _∙_ = ∀ x y → ((x ∙ x) ∙ y) ≈ (x ∙ (y ∙ y)) - RightAlternative _∙_ = ∀ x y → (x ∙ (y ∙ y)) ≈ ((x ∙ y) ∙ y) - Alternative _∙_ = (LeftAlternative _∙_ ) × (RightAlternative _∙_) - Flexible _∙_ = ∀ x y → ((x ∙ y) ∙ x) ≈ (x ∙ (y ∙ x)) - Medial _∙_ = ∀ x y u z → ((x ∙ y) ∙ (u ∙ z)) ≈ ((x ∙ u) ∙ (y ∙ z)) - LeftSemimedial _∙_ = ∀ x y z → ((x ∙ x) ∙ (y ∙ z)) ≈ ((x ∙ y) ∙ (x ∙ z)) - RightSemimedial _∙_ = ∀ x y z → ((y ∙ z) ∙ (x ∙ x)) ≈ ((y ∙ x) ∙ (z ∙ x)) - Semimedial _∙_ = (LeftSemimedial _∙_) × (RightSemimedial _∙_) - LeftBol _∙_ = ∀ x y z → (x ∙ (y ∙ (x ∙ z))) ≈ ((x ∙ (y ∙ z)) ∙ z ) - RightBol _∙_ = ∀ x y z → (((z ∙ x) ∙ y) ∙ x) ≈ (z ∙ ((x ∙ y) ∙ x)) - MiddleBol _∙_ _\\_ _//_ = ∀ x y z → (x ∙ ((y ∙ z) \\ x)) ≈ ((x // z) ∙ (y \\ x)) - ``` - -* Added new functions to `Algebra.Definitions.RawSemiring`: - ```agda - _^[_]*_ : A → ℕ → A → A - _^ᵗ_ : A → ℕ → A - ``` - -* `Algebra.Properties.Magma.Divisibility` now re-exports operations - `_∣ˡ_`, `_∤ˡ_`, `_∣ʳ_`, `_∤ʳ_` from `Algebra.Definitions.Magma`. - -* Added new proofs to `Algebra.Properties.CommutativeSemigroup`: - ```agda - interchange : Interchangable _∙_ _∙_ - xy∙xx≈x∙yxx : ∀ x y → (x ∙ y) ∙ (x ∙ x) ≈ x ∙ (y ∙ (x ∙ x)) - leftSemimedial : LeftSemimedial _∙_ - rightSemimedial : RightSemimedial _∙_ - middleSemimedial : ∀ x y z → (x ∙ y) ∙ (z ∙ x) ≈ (x ∙ z) ∙ (y ∙ x) - semimedial : Semimedial _∙_ - ``` -* Added new proof to `Algebra.Properties.Monoid.Mult`: - ```agda - ×-congˡ : ∀ {x} → (_× x) Preserves _≡_ ⟶ _≈_ - ``` - -* Added new proof to `Algebra.Properties.Monoid.Sum`: - ```agda - sum-init-last : ∀ {n} (t : Vector _ (suc n)) → sum t ≈ sum (init t) + last t - ``` - -* Added new proofs to `Algebra.Properties.Semigroup`: - ```agda - leftAlternative : LeftAlternative _∙_ - rightAlternative : RightAlternative _∙_ - alternative : Alternative _∙_ - flexible : Flexible _∙_ - ``` - -* Added new proofs to `Algebra.Properties.Semiring.Exp`: - ```agda - ^-congʳ : (x ^_) Preserves _≡_ ⟶ _≈_ - y*x^m*y^n≈x^m*y^[n+1] : (x * y ≈ y * x) → y * (x ^ m * y ^ n) ≈ x ^ m * y ^ suc n - ``` - -* Added new proofs to `Algebra.Properties.Semiring.Mult`: - ```agda - 1×-identityʳ : 1 × x ≈ x - ×-comm-* : x * (n × y) ≈ n × (x * y) - ×-assoc-* : (n × x) * y ≈ n × (x * y) - ``` - -* Added new proofs to `Algebra.Properties.Ring`: - ```agda - -1*x≈-x : ∀ x → - 1# * x ≈ - x - x+x≈x⇒x≈0 : ∀ x → x + x ≈ x → x ≈ 0# - x[y-z]≈xy-xz : ∀ x y z → x * (y - z) ≈ x * y - x * z - [y-z]x≈yx-zx : ∀ x y z → (y - z) * x ≈ (y * x) - (z * x) - ``` - -* Added new definitions to `Algebra.Structures`: - ```agda - record IsUnitalMagma (_∙_ : Op₂ A) (ε : A) : Set (a ⊔ ℓ) - record IsInvertibleMagma (_∙_ : Op₂ A) (ε : A) (_⁻¹ : Op₁ A) : Set (a ⊔ ℓ) - record IsInvertibleUnitalMagma (_∙_ : Op₂ A) (ε : A) (⁻¹ : Op₁ A) : Set (a ⊔ ℓ) - record IsQuasigroup (∙ \\ // : Op₂ A) : Set (a ⊔ ℓ) - record IsLoop (∙ \\ // : Op₂ A) (ε : A) : Set (a ⊔ ℓ) - record IsRingWithoutOne (+ * : Op₂ A) (-_ : Op₁ A) (0# : A) : Set (a ⊔ ℓ) - record IsIdempotentSemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) - record IsKleeneAlgebra (+ * : Op₂ A) (⋆ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) - record IsQuasiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where - record IsNearring (+ * : Op₂ A) (0# 1# : A) (_⁻¹ : Op₁ A) : Set (a ⊔ ℓ) where - record IsIdempotentMagma (∙ : Op₂ A) : Set (a ⊔ ℓ) - record IsAlternativeMagma (∙ : Op₂ A) : Set (a ⊔ ℓ) - record IsFlexibleMagma (∙ : Op₂ A) : Set (a ⊔ ℓ) - record IsMedialMagma (∙ : Op₂ A) : Set (a ⊔ ℓ) - record IsSemimedialMagma (∙ : Op₂ A) : Set (a ⊔ ℓ) - record IsLeftBolLoop (∙ \\ // : Op₂ A) (ε : A) : Set (a ⊔ ℓ) - record IsRightBolLoop (∙ \\ // : Op₂ A) (ε : A) : Set (a ⊔ ℓ) - record IsMoufangLoop (∙ \\ // : Op₂ A) (ε : A) : Set (a ⊔ ℓ) - record IsNonAssociativeRing (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) - record IsMiddleBolLoop (∙ \\ // : Op₂ A) (ε : A) : Set (a ⊔ ℓ) - ``` - and the existing record `IsLattice` now provides - ``` - ∨-isCommutativeSemigroup : IsCommutativeSemigroup ∨ - ∧-isCommutativeSemigroup : IsCommutativeSemigroup ∧ - ``` - and their corresponding algebraic substructures. - -* Added new records to `Algebra.Morphism.Structures`: - ```agda - record IsQuasigroupHomomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsQuasigroupMonomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsQuasigroupIsomorphism (⟦_⟧ : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) - record IsLoopHomomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsLoopMonomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsLoopIsomorphism (⟦_⟧ : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) - record IsRingWithoutOneHomomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsRingWithoutOneMonomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsRingWithoutOneIsoMorphism (⟦_⟧ : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) - record IsKleeneAlgebraHomomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsKleeneAlgebraMonomorphism (⟦_⟧ : A → B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) - record IsKleeneAlgebraIsomorphism (⟦_⟧ : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) - ``` - -* Added new proofs in `Data.Bool.Properties`: - ```agda - <-wellFounded : WellFounded _<_ - ∨-conicalˡ : LeftConical false _∨_ - ∨-conicalʳ : RightConical false _∨_ - ∨-conical : Conical false _∨_ - ∧-conicalˡ : LeftConical true _∧_ - ∧-conicalʳ : RightConical true _∧_ - ∧-conical : Conical true _∧_ - - true-xor : true xor x ≡ not x - xor-same : x xor x ≡ false - not-distribˡ-xor : not (x xor y) ≡ (not x) xor y - not-distribʳ-xor : not (x xor y) ≡ x xor (not y) - xor-assoc : Associative _xor_ - xor-comm : Commutative _xor_ - xor-identityˡ : LeftIdentity false _xor_ - xor-identityʳ : RightIdentity false _xor_ - xor-identity : Identity false _xor_ - xor-inverseˡ : LeftInverse true not _xor_ - xor-inverseʳ : RightInverse true not _xor_ - xor-inverse : Inverse true not _xor_ - ∧-distribˡ-xor : _∧_ DistributesOverˡ _xor_ - ∧-distribʳ-xor : _∧_ DistributesOverʳ _xor_ - ∧-distrib-xor : _∧_ DistributesOver _xor_ - xor-annihilates-not : (not x) xor (not y) ≡ x xor y - ``` - -* Added new functions in `Data.Fin.Base`: - ``` - finToFun : Fin (m ^ n) → (Fin n → Fin m) - funToFin : (Fin m → Fin n) → Fin (n ^ m) - quotient : Fin (m * n) → Fin m - remainder : Fin (m * n) → Fin n - ``` - -* Added new proofs in `Data.Fin.Induction`: - every (strict) partial order is well-founded and Noetherian. - - ```agda - spo-wellFounded : ∀ {r} {_⊏_ : Rel (Fin n) r} → IsStrictPartialOrder _≈_ _⊏_ → WellFounded _⊏_ - spo-noetherian : ∀ {r} {_⊏_ : Rel (Fin n) r} → IsStrictPartialOrder _≈_ _⊏_ → WellFounded (flip _⊏_) - ``` - -* Added new definitions and proofs in `Data.Fin.Permutation`: - ```agda - insert : Fin (suc m) → Fin (suc n) → Permutation m n → Permutation (suc m) (suc n) - insert-punchIn : insert i j π ⟨$⟩ʳ punchIn i k ≡ punchIn j (π ⟨$⟩ʳ k) - insert-remove : insert i (π ⟨$⟩ʳ i) (remove i π) ≈ π - remove-insert : remove i (insert i j π) ≈ π - ``` - -* In `Data.Fin.Properties`: - the proof that an injection from a type `A` into `Fin n` induces a - decision procedure for `_≡_` on `A` has been generalized to other - equivalences over `A` (i.e. to arbitrary setoids), and renamed from - `eq?` to the more descriptive `inj⇒≟` and `inj⇒decSetoid`. - -* Added new proofs in `Data.Fin.Properties`: - ``` - 1↔⊤ : Fin 1 ↔ ⊤ - - 0≢1+n : zero ≢ suc i - - ↑ˡ-injective : i ↑ˡ n ≡ j ↑ˡ n → i ≡ j - ↑ʳ-injective : n ↑ʳ i ≡ n ↑ʳ j → i ≡ j - finTofun-funToFin : funToFin ∘ finToFun ≗ id - funTofin-funToFun : finToFun (funToFin f) ≗ f - ^↔→ : Extensionality _ _ → Fin (m ^ n) ↔ (Fin n → Fin m) - - toℕ-mono-< : i < j → toℕ i ℕ.< toℕ j - toℕ-mono-≤ : i ≤ j → toℕ i ℕ.≤ toℕ j - toℕ-cancel-≤ : toℕ i ℕ.≤ toℕ j → i ≤ j - toℕ-cancel-< : toℕ i ℕ.< toℕ j → i < j - - splitAt⁻¹-↑ˡ : splitAt m {n} i ≡ inj₁ j → j ↑ˡ n ≡ i - splitAt⁻¹-↑ʳ : splitAt m {n} i ≡ inj₂ j → m ↑ʳ j ≡ i - - toℕ-combine : toℕ (combine i j) ≡ k ℕ.* toℕ i ℕ.+ toℕ j - combine-injectiveˡ : combine i j ≡ combine k l → i ≡ k - combine-injectiveʳ : combine i j ≡ combine k l → j ≡ l - combine-injective : combine i j ≡ combine k l → i ≡ k × j ≡ l - combine-surjective : ∀ i → ∃₂ λ j k → combine j k ≡ i - combine-monoˡ-< : i < j → combine i k < combine j l - - ℕ-ℕ≡toℕ‿ℕ- : n ℕ-ℕ i ≡ toℕ (n ℕ- i) - - lower₁-injective : lower₁ i n≢i ≡ lower₁ j n≢j → i ≡ j - pinch-injective : suc i ≢ j → suc i ≢ k → pinch i j ≡ pinch i k → j ≡ k - - i<1+i : i < suc i - - injective⇒≤ : ∀ {f : Fin m → Fin n} → Injective f → m ℕ.≤ n - <⇒notInjective : ∀ {f : Fin m → Fin n} → n ℕ.< m → ¬ (Injective f) - ℕ→Fin-notInjective : ∀ (f : ℕ → Fin n) → ¬ (Injective f) - cantor-schröder-bernstein : ∀ {f : Fin m → Fin n} {g : Fin n → Fin m} → Injective f → Injective g → m ≡ n - - cast-is-id : cast eq k ≡ k - subst-is-cast : subst Fin eq k ≡ cast eq k - cast-trans : cast eq₂ (cast eq₁ k) ≡ cast (trans eq₁ eq₂) k - - fromℕ≢inject₁ : {i : Fin n} → fromℕ n ≢ inject₁ i - - inject≤-trans : inject≤ (inject≤ i m≤n) n≤o ≡ inject≤ i (≤-trans m≤n n≤o) - inject≤-irrelevant : inject≤ i m≤n ≡ inject≤ i m≤n′ - ``` - -* Changed the fixity of `Data.Fin.Substitution.TermSubst._/Var_`. - ```agda - infix 8 ↦ infixl 8 - ``` - -* Added new lemmas in `Data.Fin.Substitution.Lemmas.TermLemmas`: - ``` - map-var≡ : {ρ₁ : Sub Fin m n} {ρ₂ : Sub T m n} {f : Fin m → Fin n} → - (∀ x → lookup ρ₁ x ≡ f x) → - (∀ x → lookup ρ₂ x ≡ T.var (f x)) → - map T.var ρ₁ ≡ ρ₂ - wk≡wk : map T.var VarSubst.wk ≡ T.wk {n = n} - id≡id : map T.var VarSubst.id ≡ T.id {n = n} - sub≡sub : {x : Fin n} → map T.var (VarSubst.sub x) ≡ T.sub (T.var x) - ↑≡↑ : {ρ : Sub Fin m n} → map T.var (ρ VarSubst.↑) ≡ map T.var ρ T.↑ - /Var≡/ : {ρ : Sub Fin m n} {t} → t /Var ρ ≡ t T./ map T.var ρ - sub-renaming-commutes : {ρ : Sub T m n} → - t /Var VarSubst.sub x T./ ρ ≡ t T./ ρ T.↑ T./ T.sub (lookup ρ x) - sub-commutes-with-renaming : {ρ : Sub Fin m n} → - t T./ T.sub t′ /Var ρ ≡ t /Var ρ VarSubst.↑ T./ T.sub (t′ /Var ρ) - -* Added new functions in `Data.Integer.Base`: - ``` - _^_ : ℤ → ℕ → ℤ - - +-0-rawGroup : Rawgroup 0ℓ 0ℓ - - *-rawMagma : RawMagma 0ℓ 0ℓ - *-1-rawMonoid : RawMonoid 0ℓ 0ℓ - ``` - -* Added new proofs in `Data.Integer.Properties`: - ```agda - sign-cong′ : s₁ ◃ n₁ ≡ s₂ ◃ n₂ → s₁ ≡ s₂ ⊎ (n₁ ≡ 0 × n₂ ≡ 0) - ≤-⊖ : m ℕ.≤ n → n ⊖ m ≡ + (n ∸ m) - ∣⊖∣-≤ : m ℕ.≤ n → ∣ m ⊖ n ∣ ≡ n ∸ m - ∣-∣-≤ : i ≤ j → + ∣ i - j ∣ ≡ j - i - - i^n≡0⇒i≡0 : i ^ n ≡ 0ℤ → i ≡ 0ℤ - ^-identityʳ : i ^ 1 ≡ i - ^-zeroˡ : 1 ^ n ≡ 1 - ^-*-assoc : (i ^ m) ^ n ≡ i ^ (m ℕ.* n) - ^-distribˡ-+-* : i ^ (m ℕ.+ n) ≡ i ^ m * i ^ n - - ^-isMagmaHomomorphism : IsMagmaHomomorphism ℕ.+-rawMagma *-rawMagma (i ^_) - ^-isMonoidHomomorphism : IsMonoidHomomorphism ℕ.+-0-rawMonoid *-1-rawMonoid (i ^_) - ``` - -* Added new proofs in `Data.Integer.GCD`: - ```agda - gcd-assoc : Associative gcd - gcd-zeroˡ : LeftZero 1ℤ gcd - gcd-zeroʳ : RightZero 1ℤ gcd - gcd-zero : Zero 1ℤ gcd - ``` - -* Added new functions in `Data.List`: - ```agda - takeWhileᵇ : (A → Bool) → List A → List A - dropWhileᵇ : (A → Bool) → List A → List A - filterᵇ : (A → Bool) → List A → List A - partitionᵇ : (A → Bool) → List A → List A × List A - spanᵇ : (A → Bool) → List A → List A × List A - breakᵇ : (A → Bool) → List A → List A × List A - linesByᵇ : (A → Bool) → List A → List (List A) - wordsByᵇ : (A → Bool) → List A → List (List A) - derunᵇ : (A → A → Bool) → List A → List A - deduplicateᵇ : (A → A → Bool) → List A → List A - - findᵇ : (A → Bool) → List A -> Maybe A - findIndexᵇ : (A → Bool) → (xs : List A) → Maybe $ Fin (length xs) - findIndicesᵇ : (A → Bool) → (xs : List A) → List $ Fin (length xs) - find : Decidable P → List A → Maybe A - findIndex : Decidable P → (xs : List A) → Maybe $ Fin (length xs) - findIndices : Decidable P → (xs : List A) → List $ Fin (length xs) - ``` - -* Added new functions and definitions to `Data.List.Base`: - ```agda - catMaybes : List (Maybe A) → List A - ap : List (A → B) → List A → List B - ++-rawMagma : Set a → RawMagma a _ - ++-[]-rawMonoid : Set a → RawMonoid a _ - ``` - -* Added new proofs in `Data.List.Relation.Binary.Lex.Strict`: - ```agda - xs≮[] : ¬ xs < [] - ``` - -* Added new proofs to `Data.List.Relation.Binary.Permutation.Propositional.Properties`: - ```agda - Any-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) → (ix : Any P xs) → Any-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix - ∈-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) → (ix : x ∈ xs) → ∈-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix - ``` - -* Added new functions in `Data.List.Relation.Unary.All`: - ``` - decide : Π[ P ∪ Q ] → Π[ All P ∪ Any Q ] - ``` - -* Added new functions in `Data.List.Fresh.Relation.Unary.All`: - ``` - decide : Π[ P ∪ Q ] → Π[ All {R = R} P ∪ Any Q ] - ``` - -* Added new proofs to `Data.List.Membership.Propositional.Properties`: - ```agda - mapWith∈-id : mapWith∈ xs (λ {x} _ → x) ≡ xs - map-mapWith∈ : map g (mapWith∈ xs f) ≡ mapWith∈ xs (g ∘′ f) - ``` - -* Added new proofs to `Data.List.Membership.Setoid.Properties`: - ```agda - mapWith∈-id : mapWith∈ xs (λ {x} _ → x) ≡ xs - map-mapWith∈ : map g (mapWith∈ xs f) ≡ mapWith∈ xs (g ∘′ f) - index-injective : index x₁∈xs ≡ index x₂∈xs → x₁ ≈ x₂ - ``` - -* Add new proofs in `Data.List.Properties`: - ```agda - ∈⇒∣product : n ∈ ns → n ∣ product ns - ∷ʳ-++ : xs ∷ʳ a ++ ys ≡ xs ++ a ∷ ys - - concatMap-cong : f ≗ g → concatMap f ≗ concatMap g - concatMap-pure : concatMap [_] ≗ id - concatMap-map : concatMap g (map f xs) ≡ concatMap (g ∘′ f) xs - map-concatMap : map f ∘′ concatMap g ≗ concatMap (map f ∘′ g) - - length-isMagmaHomomorphism : (A : Set a) → IsMagmaHomomorphism (++-rawMagma A) +-rawMagma length - length-isMonoidHomomorphism : (A : Set a) → IsMonoidHomomorphism (++-[]-rawMonoid A) +-0-rawMonoid length - - take-map : take n (map f xs) ≡ map f (take n xs) - drop-map : drop n (map f xs) ≡ map f (drop n xs) - head-map : head (map f xs) ≡ Maybe.map f (head xs) - - take-suc : take (suc m) xs ≡ take m xs ∷ʳ lookup xs i - take-suc-tabulate : take (suc m) (tabulate f) ≡ take m (tabulate f) ∷ʳ f i - drop-take-suc : drop m (take (suc m) xs) ≡ [ lookup xs i ] - drop-take-suc-tabulate : drop m (take (suc m) (tabulate f)) ≡ [ f i ] - - take-all : n ≥ length xs → take n xs ≡ xs - - take-[] : take m [] ≡ [] - drop-[] : drop m [] ≡ [] - - map-replicate : map f (replicate n x) ≡ replicate n (f x) - - drop-drop : drop n (drop m xs) ≡ drop (m + n) xs - ``` - -* Added new patterns and definitions to `Data.Nat.Base`: - ```agda - pattern z0⇒n≢0 : n > 0 → n ≢ 0 - m^n≢0 : .{{_ : NonZero m}} → NonZero (m ^ n) - m*n≢0 : .{{_ : NonZero m}} .{{_ : NonZero n}} → NonZero (m * n) - m≤n⇒n∸m≤n : m ≤ n → n ∸ m ≤ n - - ≤-pred : suc m ≤ suc n → m ≤ n - s0 : ∀ m .{{_ : NonZero m}} n → m ^ n > 0 - - ^-monoˡ-≤ : ∀ n → (_^ n) Preserves _≤_ ⟶ _≤_ - ^-monoʳ-≤ : ∀ m .{{_ : NonZero m}} → (m ^_) Preserves _≤_ ⟶ _≤_ - ^-monoˡ-< : ∀ n .{{_ : NonZero n}} → (_^ n) Preserves _<_ ⟶ _<_ - ^-monoʳ-< : ∀ m → 1 < m → (m ^_) Preserves _<_ ⟶ _<_ - - n≡⌊n+n/2⌋ : n ≡ ⌊ n + n /2⌋ - n≡⌈n+n/2⌉ : n ≡ ⌈ n + n /2⌉ - - m-connex : Connex _≥_ _>_ - <-≤-connex : Connex _<_ _≤_ - >-≥-connex : Connex _>_ _≥_ - <-cmp : Trichotomous _≡_ _<_ - anyUpTo? : (P? : Decidable P) → ∀ v → Dec (∃ λ n → n < v × P n) - allUpTo? : (P? : Decidable P) → ∀ v → Dec (∀ {n} → n < v → P n) - ``` - -* Added new proofs in `Data.Nat.Combinatorics`: - ```agda - [n-k]*[n-k-1]!≡[n-k]! : k < n → (n ∸ k) * (n ∸ suc k) ! ≡ (n ∸ k) ! - [n-k]*d[k+1]≡[k+1]*d[k] : k < n → (n ∸ k) * ((suc k) ! * (n ∸ suc k) !) ≡ (suc k) * (k ! * (n ∸ k) !) - k![n∸k]!∣n! : k ≤ n → k ! * (n ∸ k) ! ∣ n ! - nP1≡n : n P 1 ≡ n - nC1≡n : n C 1 ≡ n - nCk+nC[k+1]≡[n+1]C[k+1] : n C k + n C (suc k) ≡ suc n C suc k - ``` - -* Added new proofs in `Data.Nat.DivMod`: - ```agda - m%n≤n : .{{_ : NonZero n}} → m % n ≤ n - m*n/m!≡n/[m∸1]! : .{{_ : NonZero m}} → m * n / m ! ≡ n / (pred m) ! - - %-congˡ : .⦃ _ : NonZero o ⦄ → m ≡ n → m % o ≡ n % o - %-congʳ : .⦃ _ : NonZero m ⦄ .⦃ _ : NonZero n ⦄ → m ≡ n → o % m ≡ o % n - m≤n⇒[n∸m]%m≡n%m : .⦃ _ : NonZero m ⦄ → m ≤ n → (n ∸ m) % m ≡ n % m - m*n≤o⇒[o∸m*n]%n≡o%n : .⦃ _ : NonZero n ⦄ → m * n ≤ o → (o ∸ m * n) % n ≡ o % n - m∣n⇒o%n%m≡o%m : .⦃ _ : NonZero m ⦄ .⦃ _ : NonZero n ⦄ → m ∣ n → o % n % m ≡ o % m - m?_ : Decidable _>_ - - +-*-rawNearSemiring : RawNearSemiring 0ℓ 0ℓ - +-*-rawSemiring : RawSemiring 0ℓ 0ℓ - toℚᵘ-isNearSemiringHomomorphism-+-* : IsNearSemiringHomomorphism +-*-rawNearSemiring ℚᵘ.+-*-rawNearSemiring toℚᵘ - toℚᵘ-isNearSemiringMonomorphism-+-* : IsNearSemiringMonomorphism +-*-rawNearSemiring ℚᵘ.+-*-rawNearSemiring toℚᵘ - toℚᵘ-isSemiringHomomorphism-+-* : IsSemiringHomomorphism +-*-rawSemiring ℚᵘ.+-*-rawSemiring toℚᵘ - toℚᵘ-isSemiringMonomorphism-+-* : IsSemiringMonomorphism +-*-rawSemiring ℚᵘ.+-*-rawSemiring toℚᵘ - - pos⇒nonZero : .{{Positive p}} → NonZero p - neg⇒nonZero : .{{Negative p}} → NonZero p - nonZero⇒1/nonZero : .{{_ : NonZero p}} → NonZero (1/ p) - ``` - -* Added new rounding functions in `Data.Rational.Unnormalised.Base`: - ```agda - floor ceiling truncate round ⌊_⌋ ⌈_⌉ [_] : ℚᵘ → ℤ - fracPart : ℚᵘ → ℚᵘ - ``` - -* Added new definitions in `Data.Rational.Unnormalised.Properties`: - ```agda - +-*-rawNearSemiring : RawNearSemiring 0ℓ 0ℓ - +-*-rawSemiring : RawSemiring 0ℓ 0ℓ - - ≰⇒≥ : _≰_ ⇒ _≥_ - - _≥?_ : Decidable _≥_ - _>?_ : Decidable _>_ - - *-mono-≤-nonNeg : .{{_ : NonNegative p}} .{{_ : NonNegative r}} → p ≤ q → r ≤ s → p * r ≤ q * s - *-mono-<-nonNeg : .{{_ : NonNegative p}} .{{_ : NonNegative r}} → p < q → r < s → p * r < q * s - 1/-antimono-≤-pos : .{{_ : Positive p}} .{{_ : Positive q}} → p ≤ q → 1/ q ≤ 1/ p - ⊓-mono-< : _⊓_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_ - ⊔-mono-< : _⊔_ Preserves₂ _<_ ⟶ _<_ ⟶ _<_ - - pos⇒nonZero : .{{_ : Positive p}} → NonZero p - neg⇒nonZero : .{{_ : Negative p}} → NonZero p - pos+pos⇒pos : .{{_ : Positive p}} .{{_ : Positive q}} → Positive (p + q) - nonNeg+nonNeg⇒nonNeg : .{{_ : NonNegative p}} .{{_ : NonNegative q}} → NonNegative (p + q) - pos*pos⇒pos : .{{_ : Positive p}} .{{_ : Positive q}} → Positive (p * q) - nonNeg*nonNeg⇒nonNeg : .{{_ : NonNegative p}} .{{_ : NonNegative q}} → NonNegative (p * q) - pos⊓pos⇒pos : .{{_ : Positive p}} .{{_ : Positive q}} → Positive (p ⊓ q) - pos⊔pos⇒pos : .{{_ : Positive p}} .{{_ : Positive q}} → Positive (p ⊔ q) - 1/nonZero⇒nonZero : .{{_ : NonZero p}} → NonZero (1/ p) - ``` - -* Added new functions to `Data.Product.Nary.NonDependent`: - ```agda - zipWith : (∀ k → Projₙ as k → Projₙ bs k → Projₙ cs k) → - Product n as → Product n bs → Product n cs - ``` - -* Added new proof to `Data.Product.Properties`: - ```agda - map-cong : f ≗ g → h ≗ i → map f h ≗ map g i - ``` - -* Added new definitions to `Data.Product.Properties` and - `Function.Related.TypeIsomorphisms` for convenience: - ``` - Σ-≡,≡→≡ : (∃ λ (p : proj₁ p₁ ≡ proj₁ p₂) → subst B p (proj₂ p₁) ≡ proj₂ p₂) → p₁ ≡ p₂ - Σ-≡,≡←≡ : p₁ ≡ p₂ → (∃ λ (p : proj₁ p₁ ≡ proj₁ p₂) → subst B p (proj₂ p₁) ≡ proj₂ p₂) - ×-≡,≡→≡ : (proj₁ p₁ ≡ proj₁ p₂ × proj₂ p₁ ≡ proj₂ p₂) → p₁ ≡ p₂ - ×-≡,≡←≡ : p₁ ≡ p₂ → (proj₁ p₁ ≡ proj₁ p₂ × proj₂ p₁ ≡ proj₂ p₂) - ``` - -* Added new proofs to `Data.Product.Relation.Binary.Lex.Strict` - ```agda - ×-respectsʳ : Transitive _≈₁_ → - _<₁_ Respectsʳ _≈₁_ → _<₂_ Respectsʳ _≈₂_ → _<ₗₑₓ_ Respectsʳ _≋_ - ×-respectsˡ : Symmetric _≈₁_ → Transitive _≈₁_ → - _<₁_ Respectsˡ _≈₁_ → _<₂_ Respectsˡ _≈₂_ → _<ₗₑₓ_ Respectsˡ _≋_ - ×-wellFounded' : Transitive _≈₁_ → _<₁_ Respectsʳ _≈₁_ → - WellFounded _<₁_ → WellFounded _<₂_ → WellFounded _<ₗₑₓ_ - ``` - -* Added new definitions to `Data.Sign.Base`: - ```agda - *-rawMagma : RawMagma 0ℓ 0ℓ - *-1-rawMonoid : RawMonoid 0ℓ 0ℓ - *-1-rawGroup : RawGroup 0ℓ 0ℓ - ``` - -* Added new proofs to `Data.Sign.Properties`: - ```agda - *-inverse : Inverse + id _*_ - *-isCommutativeSemigroup : IsCommutativeSemigroup _*_ - *-isCommutativeMonoid : IsCommutativeMonoid _*_ + - *-isGroup : IsGroup _*_ + id - *-isAbelianGroup : IsAbelianGroup _*_ + id - *-commutativeSemigroup : CommutativeSemigroup 0ℓ 0ℓ - *-commutativeMonoid : CommutativeMonoid 0ℓ 0ℓ - *-group : Group 0ℓ 0ℓ - *-abelianGroup : AbelianGroup 0ℓ 0ℓ - ≡-isDecEquivalence : IsDecEquivalence _≡_ - ``` - -* Added new functions in `Data.String.Base`: - ```agda - wordsByᵇ : (Char → Bool) → String → List String - linesByᵇ : (Char → Bool) → String → List String - ``` - -* Added new proofs in `Data.String.Properties`: - ``` - ≤-isDecTotalOrder-≈ : IsDecTotalOrder _≈_ _≤_ - ≤-decTotalOrder-≈ : DecTotalOrder _ _ _ - ``` -* Added new definitions in `Data.Sum.Properties`: - ``` - swap-↔ : (A ⊎ B) ↔ (B ⊎ A) - ``` - -* Added new proofs in `Data.Sum.Properties`: - ``` - map-assocˡ : (f : A → C) (g : B → D) (h : C → F) → - map (map f g) h ∘ assocˡ ≗ assocˡ ∘ map f (map g h) - map-assocʳ : (f : A → C) (g : B → D) (h : C → F) → - map f (map g h) ∘ assocʳ ≗ assocʳ ∘ map (map f g) h - ``` - -* Made `Map` public in `Data.Tree.AVL.IndexedMap` - -* Added new definitions in `Data.Vec.Base`: - ```agda - truncate : m ≤ n → Vec A n → Vec A m - pad : m ≤ n → A → Vec A m → Vec A n - - FoldrOp A B = ∀ {n} → A → B n → B (suc n) - FoldlOp A B = ∀ {n} → B n → A → B (suc n) - - foldr′ : (A → B → B) → B → Vec A n → B - foldl′ : (B → A → B) → B → Vec A n → B - countᵇ : (A → Bool) → Vec A n → ℕ - - iterate : (A → A) → A → Vec A n - - diagonal : Vec (Vec A n) n → Vec A n - DiagonalBind._>>=_ : Vec A n → (A → Vec B n) → Vec B n - - _ʳ++_ : Vec A m → Vec A n → Vec A (m + n) - - cast : .(eq : m ≡ n) → Vec A m → Vec A n - ``` - -* Added new instance in `Data.Vec.Effectful`: - ```agda - monad : RawMonad (λ (A : Set a) → Vec A n) - ``` - -* Added new proofs in `Data.Vec.Properties`: - ```agda - padRight-refl : padRight ≤-refl a xs ≡ xs - padRight-replicate : replicate a ≡ padRight le a (replicate a) - padRight-trans : padRight (≤-trans m≤n n≤p) a xs ≡ padRight n≤p a (padRight m≤n a xs) - - truncate-refl : truncate ≤-refl xs ≡ xs - truncate-trans : truncate (≤-trans m≤n n≤p) xs ≡ truncate m≤n (truncate n≤p xs) - truncate-padRight : truncate m≤n (padRight m≤n a xs) ≡ xs - - map-const : map (const x) xs ≡ replicate x - map-⊛ : map f xs ⊛ map g xs ≡ map (f ˢ g) xs - map-++ : map f (xs ++ ys) ≡ map f xs ++ map f ys - map-is-foldr : map f ≗ foldr (Vec B) (λ x ys → f x ∷ ys) [] - map-∷ʳ : map f (xs ∷ʳ x) ≡ (map f xs) ∷ʳ (f x) - map-reverse : map f (reverse xs) ≡ reverse (map f xs) - map-ʳ++ : map f (xs ʳ++ ys) ≡ map f xs ʳ++ map f ys - map-insert : map f (insert xs i x) ≡ insert (map f xs) i (f x) - toList-map : toList (map f xs) ≡ List.map f (toList xs) - - lookup-concat : lookup (concat xss) (combine i j) ≡ lookup (lookup xss i) j - - ⊛-is->>= : fs ⊛ xs ≡ fs >>= flip map xs - lookup-⊛* : lookup (fs ⊛* xs) (combine i j) ≡ (lookup fs i $ lookup xs j) - ++-is-foldr : xs ++ ys ≡ foldr ((Vec A) ∘ (_+ n)) _∷_ ys xs - []≔-++-↑ʳ : (xs ++ ys) [ m ↑ʳ i ]≔ y ≡ xs ++ (ys [ i ]≔ y) - unfold-ʳ++ : xs ʳ++ ys ≡ reverse xs ++ ys - - foldl-universal : ∀ (h : ∀ {c} (C : ℕ → Set c) (g : FoldlOp A C) (e : C zero) → ∀ {n} → Vec A n → C n) → - (∀ ... → h C g e [] ≡ e) → - (∀ ... → h C g e ∘ (x ∷_) ≗ h (C ∘ suc) g (g e x)) → - h B f e ≗ foldl B f e - foldl-fusion : h d ≡ e → (∀ ... → h (f b x) ≡ g (h b) x) → h ∘ foldl B f d ≗ foldl C g e - foldl-∷ʳ : foldl B f e (ys ∷ʳ y) ≡ f (foldl B f e ys) y - foldl-[] : foldl B f e [] ≡ e - foldl-reverse : foldl B {n} f e ∘ reverse ≗ foldr B (flip f) e - - foldr-[] : foldr B f e [] ≡ e - foldr-++ : foldr B f e (xs ++ ys) ≡ foldr (B ∘ (_+ n)) f (foldr B f e ys) xs - foldr-∷ʳ : foldr B f e (ys ∷ʳ y) ≡ foldr (B ∘ suc) f (f y e) ys - foldr-ʳ++ : foldr B f e (xs ʳ++ ys) ≡ foldl (B ∘ (_+ n)) (flip f) (foldr B f e ys) xs - foldr-reverse : foldr B f e ∘ reverse ≗ foldl B (flip f) e - - ∷ʳ-injective : xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys × x ≡ y - ∷ʳ-injectiveˡ : xs ∷ʳ x ≡ ys ∷ʳ y → xs ≡ ys - ∷ʳ-injectiveʳ : xs ∷ʳ x ≡ ys ∷ʳ y → x ≡ y - - unfold-∷ʳ : cast eq (xs ∷ʳ x) ≡ xs ++ [ x ] - init-∷ʳ : init (xs ∷ʳ x) ≡ xs - last-∷ʳ : last (xs ∷ʳ x) ≡ x - cast-∷ʳ : cast eq (xs ∷ʳ x) ≡ (cast (cong pred eq) xs) ∷ʳ x - ++-∷ʳ : cast eq ((xs ++ ys) ∷ʳ z) ≡ xs ++ (ys ∷ʳ z) - - reverse-∷ : reverse (x ∷ xs) ≡ reverse xs ∷ʳ x - reverse-involutive : Involutive _≡_ reverse - reverse-reverse : reverse xs ≡ ys → reverse ys ≡ xs - reverse-injective : reverse xs ≡ reverse ys → xs ≡ ys - - transpose-replicate : transpose (replicate xs) ≡ map replicate xs - toList-replicate : toList (replicate {n = n} a) ≡ List.replicate n a - - toList-++ : toList (xs ++ ys) ≡ toList xs List.++ toList ys - - cast-is-id : cast eq xs ≡ xs - subst-is-cast : subst (Vec A) eq xs ≡ cast eq xs - cast-sym : cast eq xs ≡ ys → cast (sym eq) ys ≡ xs - cast-trans : cast eq₂ (cast eq₁ xs) ≡ cast (trans eq₁ eq₂) xs - map-cast : map f (cast eq xs) ≡ cast eq (map f xs) - lookup-cast : lookup (cast eq xs) (Fin.cast eq i) ≡ lookup xs i - lookup-cast₁ : lookup (cast eq xs) i ≡ lookup xs (Fin.cast (sym eq) i) - lookup-cast₂ : lookup xs (Fin.cast eq i) ≡ lookup (cast (sym eq) xs) i - cast-reverse : cast eq ∘ reverse ≗ reverse ∘ cast eq - - zipwith-++ : zipWith f (xs ++ ys) (xs' ++ ys') ≡ zipWith f xs xs' ++ zipWith f ys ys' - - ++-assoc : cast eq ((xs ++ ys) ++ zs) ≡ xs ++ (ys ++ zs) - ++-identityʳ : cast eq (xs ++ []) ≡ xs - init-reverse : init ∘ reverse ≗ reverse ∘ tail - last-reverse : last ∘ reverse ≗ head - reverse-++ : cast eq (reverse (xs ++ ys)) ≡ reverse ys ++ reverse xs - - toList-cast : toList (cast eq xs) ≡ toList xs - cast-fromList : cast _ (fromList xs) ≡ fromList ys - fromList-map : cast _ (fromList (List.map f xs)) ≡ map f (fromList xs) - fromList-++ : cast _ (fromList (xs List.++ ys)) ≡ fromList xs ++ fromList ys - - truncate≡take : .(eq : n ≡ m + o) → truncate m≤n xs ≡ take m (cast eq xs) - take≡truncate : take m xs ≡ truncate (m≤m+n m n) xs - lookup-truncate : lookup (truncate m≤n xs) i ≡ lookup xs (Fin.inject≤ i m≤n) - lookup-take-inject≤ : lookup (take m xs) i ≡ lookup xs (Fin.inject≤ i (m≤m+n m n)) - ``` - -* Added new proofs in `Data.Vec.Membership.Propositional.Properties`: - ```agda - index-∈-fromList⁺ : Any.index (∈-fromList⁺ v∈xs) ≡ indexₗ v∈xs - ``` - -* Added new proofs in `Data.Vec.Functional.Properties`: - ``` - map-updateAt : f ∘ g ≗ h ∘ f → map f (updateAt i g xs) ≗ updateAt i h (map f xs) - ``` - -* Added new proofs in `Data.Vec.Relation.Binary.Lex.Strict`: - ```agda - xs≮[] : {xs : Vec A n} → ¬ xs < [] - <-respectsˡ : IsPartialEquivalence _≈_ → _≺_ Respectsˡ _≈_ → - ∀ {m n} → _Respectsˡ_ (_<_ {m} {n}) _≋_ - <-respectsʳ : IsPartialEquivalence _≈_ → _≺_ Respectsʳ _≈_ → - ∀ {m n} → _Respectsʳ_ (_<_ {m} {n}) _≋_ - <-wellFounded : Transitive _≈_ → _≺_ Respectsʳ _≈_ → WellFounded _≺_ → - ∀ {n} → WellFounded (_<_ {n}) -``` - -* Added new functions in `Data.Vec.Relation.Unary.Any`: - ``` - lookup : Any P xs → A - ``` - -* Added new functions in `Data.Vec.Relation.Unary.All`: - ``` - decide : Π[ P ∪ Q ] → Π[ All P ∪ Any Q ] - ``` - -* Added vector associativity proof to `Data.Vec.Relation.Binary.Equality.Setoid`: - ``` - ++-assoc : (xs ++ ys) ++ zs ≋ xs ++ (ys ++ zs) - ``` - -* Added new functions in `Effect.Monad.State`: - ``` - runState : State s a → s → a × s - evalState : State s a → s → a - execState : State s a → s → s - ``` - -* Added new proofs in `Function.Construct.Symmetry`: - ``` - bijective : Bijective ≈₁ ≈₂ f → Symmetric ≈₂ → Transitive ≈₂ → Congruent ≈₁ ≈₂ f → Bijective ≈₂ ≈₁ f⁻¹ - isBijection : IsBijection ≈₁ ≈₂ f → Congruent ≈₂ ≈₁ f⁻¹ → IsBijection ≈₂ ≈₁ f⁻¹ - isBijection-≡ : IsBijection ≈₁ _≡_ f → IsBijection _≡_ ≈₁ f⁻¹ - bijection : Bijection R S → Congruent IB.Eq₂._≈_ IB.Eq₁._≈_ f⁻¹ → Bijection S R - bijection-≡ : Bijection R (setoid B) → Bijection (setoid B) R - sym-⤖ : A ⤖ B → B ⤖ A - ``` - -* Added new operations in `Function.Strict`: - ``` - _!|>_ : (a : A) → (∀ a → B a) → B a - _!|>′_ : A → (A → B) → B - ``` - -* Added new definition to the `Surjection` module in `Function.Related.Surjection`: - ``` - f⁻ = proj₁ ∘ surjective - ``` - -* Added new operations in `IO`: - ``` - Colist.forM : Colist A → (A → IO B) → IO (Colist B) - Colist.forM′ : Colist A → (A → IO B) → IO ⊤ - - List.forM : List A → (A → IO B) → IO (List B) - List.forM′ : List A → (A → IO B) → IO ⊤ - ``` - -* Added new operations in `IO.Base`: - ``` - lift! : IO A → IO (Lift b A) - _<$_ : B → IO A → IO B - _=<<_ : (A → IO B) → IO A → IO B - _<<_ : IO B → IO A → IO B - lift′ : Prim.IO ⊤ → IO {a} ⊤ - - when : Bool → IO ⊤ → IO ⊤ - unless : Bool → IO ⊤ → IO ⊤ - - whenJust : Maybe A → (A → IO ⊤) → IO ⊤ - untilJust : IO (Maybe A) → IO A - untilRight : (A → IO (A ⊎ B)) → A → IO B - ``` - -* Added new functions in `Reflection.AST.Term`: - ``` - stripPis : Term → List (String × Arg Type) × Term - prependLams : List (String × Visibility) → Term → Term - prependHLams : List String → Term → Term - prependVLams : List String → Term → Term - ``` - -* Added new operations in `Relation.Binary.Construct.Closure.Equivalence`: - ``` - fold : IsEquivalence _∼_ → _⟶_ ⇒ _∼_ → EqClosure _⟶_ ⇒ _∼_ - gfold : IsEquivalence _∼_ → _⟶_ =[ f ]⇒ _∼_ → EqClosure _⟶_ =[ f ]⇒ _∼_ - return : _⟶_ ⇒ EqClosure _⟶_ - join : EqClosure (EqClosure _⟶_) ⇒ EqClosure _⟶_ - _⋆ : _⟶₁_ ⇒ EqClosure _⟶₂_ → EqClosure _⟶₁_ ⇒ EqClosure _⟶₂_ - ``` - -* Added new operations in `Relation.Binary.Construct.Closure.Symmetric`: - ``` - fold : Symmetric _∼_ → _⟶_ ⇒ _∼_ → SymClosure _⟶_ ⇒ _∼_ - gfold : Symmetric _∼_ → _⟶_ =[ f ]⇒ _∼_ → SymClosure _⟶_ =[ f ]⇒ _∼_ - return : _⟶_ ⇒ SymClosure _⟶_ - join : SymClosure (SymClosure _⟶_) ⇒ SymClosure _⟶_ - _⋆ : _⟶₁_ ⇒ SymClosure _⟶₂_ → SymClosure _⟶₁_ ⇒ SymClosure _⟶₂_ - ``` - -* Added new proofs to `Relation.Binary.Lattice.Properties.{Join,Meet}Semilattice`: - ```agda - isPosemigroup : IsPosemigroup _≈_ _≤_ _∨_ - posemigroup : Posemigroup c ℓ₁ ℓ₂ - ≈-dec⇒≤-dec : Decidable _≈_ → Decidable _≤_ - ≈-dec⇒isDecPartialOrder : Decidable _≈_ → IsDecPartialOrder _≈_ _≤_ - ``` - -* Added new proofs to `Relation.Binary.Lattice.Properties.Bounded{Join,Meet}Semilattice`: - ```agda - isCommutativePomonoid : IsCommutativePomonoid _≈_ _≤_ _∨_ ⊥ - commutativePomonoid : CommutativePomonoid c ℓ₁ ℓ₂ - ``` - -* Added new proofs to `Relation.Binary.Properties.Poset`: - ```agda - ≤-dec⇒≈-dec : Decidable _≤_ → Decidable _≈_ - ≤-dec⇒isDecPartialOrder : Decidable _≤_ → IsDecPartialOrder _≈_ _≤_ - ``` - -* Added new proofs in `Relation.Binary.Properties.StrictPartialOrder`: - ```agda - >-strictPartialOrder : StrictPartialOrder s₁ s₂ s₃ - ``` - -* Added new proofs in `Relation.Binary.PropositionalEquality.Properties`: - ``` - subst-application′ : subst Q eq (f x p) ≡ f y (subst P eq p) - sym-cong : sym (cong f p) ≡ cong f (sym p) - ``` - -* Added new proofs in `Relation.Binary.HeterogeneousEquality`: - ``` - subst₂-removable : subst₂ _∼_ eq₁ eq₂ p ≅ p - ``` - -* Added new definitions in `Relation.Unary`: - ``` - _≐_ : Pred A ℓ₁ → Pred A ℓ₂ → Set _ - _≐′_ : Pred A ℓ₁ → Pred A ℓ₂ → Set _ - _∖_ : Pred A ℓ₁ → Pred A ℓ₂ → Pred A _ - ``` - -* Added new proofs in `Relation.Unary.Properties`: - ``` - ⊆-reflexive : _≐_ ⇒ _⊆_ - ⊆-antisym : Antisymmetric _≐_ _⊆_ - ⊆-min : Min _⊆_ ∅ - ⊆-max : Max _⊆_ U - ⊂⇒⊆ : _⊂_ ⇒ _⊆_ - ⊂-trans : Trans _⊂_ _⊂_ _⊂_ - ⊂-⊆-trans : Trans _⊂_ _⊆_ _⊂_ - ⊆-⊂-trans : Trans _⊆_ _⊂_ _⊂_ - ⊂-respʳ-≐ : _⊂_ Respectsʳ _≐_ - ⊂-respˡ-≐ : _⊂_ Respectsˡ _≐_ - ⊂-resp-≐ : _⊂_ Respects₂ _≐_ - ⊂-irrefl : Irreflexive _≐_ _⊂_ - ⊂-antisym : Antisymmetric _≐_ _⊂_ - ∅-⊆′ : (P : Pred A ℓ) → ∅ ⊆′ P - ⊆′-U : (P : Pred A ℓ) → P ⊆′ U - ⊆′-refl : Reflexive {A = Pred A ℓ} _⊆′_ - ⊆′-reflexive : _≐′_ ⇒ _⊆′_ - ⊆′-trans : Trans _⊆′_ _⊆′_ _⊆′_ - ⊆′-antisym : Antisymmetric _≐′_ _⊆′_ - ⊆′-min : Min _⊆′_ ∅ - ⊆′-max : Max _⊆′_ U - ⊂′⇒⊆′ : _⊂′_ ⇒ _⊆′_ - ⊂′-trans : Trans _⊂′_ _⊂′_ _⊂′_ - ⊂′-⊆′-trans : Trans _⊂′_ _⊆′_ _⊂′_ - ⊆′-⊂′-trans : Trans _⊆′_ _⊂′_ _⊂′_ - ⊂′-respʳ-≐′ : _⊂′_ Respectsʳ _≐′_ - ⊂′-respˡ-≐′ : _⊂′_ Respectsˡ _≐′_ - ⊂′-resp-≐′ : _⊂′_ Respects₂ _≐′_ - ⊂′-irrefl : Irreflexive _≐′_ _⊂′_ - ⊂′-antisym : Antisymmetric _≐′_ _⊂′_ - ⊆⇒⊆′ : _⊆_ ⇒ _⊆′_ - ⊆′⇒⊆ : _⊆′_ ⇒ _⊆_ - ⊂⇒⊂′ : _⊂_ ⇒ _⊂′_ - ⊂′⇒⊂ : _⊂′_ ⇒ _⊂_ - ≐-refl : Reflexive _≐_ - ≐-sym : Sym _≐_ _≐_ - ≐-trans : Trans _≐_ _≐_ _≐_ - ≐′-refl : Reflexive _≐′_ - ≐′-sym : Sym _≐′_ _≐′_ - ≐′-trans : Trans _≐′_ _≐′_ _≐′_ - ≐⇒≐′ : _≐_ ⇒ _≐′_ - ≐′⇒≐ : _≐′_ ⇒ _≐_ - - U-irrelevant : Irrelevant {A = A} U - ∁-irrelevant : (P : Pred A ℓ) → Irrelevant (∁ P) - ``` - -* Generalised proofs in `Relation.Unary.Properties`: - ``` - ⊆-trans : Trans _⊆_ _⊆_ _⊆_ - ``` - -* Added new proofs in `Relation.Binary.Properties.Setoid`: - ``` - ≈-isPreorder : IsPreorder _≈_ _≈_ - ≈-isPartialOrder : IsPartialOrder _≈_ _≈_ - - ≈-preorder : Preorder a ℓ ℓ - ≈-poset : Poset a ℓ ℓ - ``` - -* Added new definitions in `Relation.Binary.Definitions`: - ``` - Cotransitive _#_ = ∀ {x y} → x # y → ∀ z → (x # z) ⊎ (z # y) - Tight _≈_ _#_ = ∀ x y → (¬ x # y → x ≈ y) × (x ≈ y → ¬ x # y) - - Monotonic₁ _≤_ _⊑_ f = f Preserves _≤_ ⟶ _⊑_ - Antitonic₁ _≤_ _⊑_ f = f Preserves (flip _≤_) ⟶ _⊑_ - Monotonic₂ _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ _≤_ ⟶ _⊑_ ⟶ _≼_ - MonotonicAntitonic _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ _≤_ ⟶ (flip _⊑_) ⟶ _≼_ - AntitonicMonotonic _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ (flip _≤_) ⟶ _⊑_ ⟶ _≼_ - Antitonic₂ _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ (flip _≤_) ⟶ (flip _⊑_) ⟶ _≼_ - Adjoint _≤_ _⊑_ f g = ∀ {x y} → (f x ⊑ y → x ≤ g y) × (x ≤ g y → f x ⊑ y) - ``` - -* Added new definitions in `Relation.Binary.Bundles`: - ``` - record ApartnessRelation c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - ``` - -* Added new definitions in `Relation.Binary.Structures`: - ``` - record IsApartnessRelation (_#_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where - ``` - -* Added new proofs to `Relation.Binary.Consequences`: - ``` - sym⇒¬-sym : Symmetric _∼_ → Symmetric _≁_ - cotrans⇒¬-trans : Cotransitive _∼_ → Transitive _≁_ - irrefl⇒¬-refl : Reflexive _≈_ → Irreflexive _≈_ _∼_ → Reflexive _≁_ - mono₂⇒cong₂ : Symmetric ≈₁ → ≈₁ ⇒ ≤₁ → Antisymmetric ≈₂ ≤₂ → ∀ {f} → - f Preserves₂ ≤₁ ⟶ ≤₁ ⟶ ≤₂ → - f Preserves₂ ≈₁ ⟶ ≈₁ ⟶ ≈₂ - ``` - -* Added new proofs to `Relation.Binary.Construct.Closure.Transitive`: - ``` - accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x - wellFounded⁻ : WellFounded _∼⁺_ → WellFounded _∼_ - accessible : ∀ {x} → Acc _∼_ x → Acc _∼⁺_ x - ``` - -* Added new operations in `Relation.Binary.PropositionalEquality.Properties`: - ``` - J : (B : (y : A) → x ≡ y → Set b) (p : x ≡ y) → B x refl → B y p - dcong : (p : x ≡ y) → subst B p (f x) ≡ f y - dcong₂ : (p : x₁ ≡ x₂) → subst B p y₁ ≡ y₂ → f x₁ y₁ ≡ f x₂ y₂ - dsubst₂ : (p : x₁ ≡ x₂) → subst B p y₁ ≡ y₂ → C x₁ y₁ → C x₂ y₂ - ddcong₂ : (p : x₁ ≡ x₂) (q : subst B p y₁ ≡ y₂) → dsubst₂ C p q (f x₁ y₁) ≡ f x₂ y₂ - - isPartialOrder : IsPartialOrder _≡_ _≡_ - poset : Set a → Poset _ _ _ - ``` - -* Added new operations in `System.Exit`: - ``` - isSuccess : ExitCode → Bool - isFailure : ExitCode → Bool - ``` - -* Added new functions in `Codata.Guarded.Stream`: - ``` - transpose : List (Stream A) → Stream (List A) - transpose⁺ : List⁺ (Stream A) → Stream (List⁺ A) - concat : Stream (List⁺ A) → Stream A - ``` - -* Added new proofs in `Codata.Guarded.Stream.Properties`: - ``` - cong-concat : {ass bss : Stream (List⁺.List⁺ A)} → ass ≈ bss → concat ass ≈ concat bss - map-concat : ∀ (f : A → B) ass → map f (concat ass) ≈ concat (map (List⁺.map f) ass) - lookup-transpose : ∀ n (ass : List (Stream A)) → lookup n (transpose ass) ≡ List.map (lookup n) ass - lookup-transpose⁺ : ∀ n (ass : List⁺ (Stream A)) → lookup n (transpose⁺ ass) ≡ List⁺.map (lookup n) ass - ``` - -* Added new corollaries in `Data.List.Membership.Setoid.Properties`: - ``` - ∈-++⁺∘++⁻ : ∀ {v} xs {ys} (p : v ∈ xs ++ ys) → [ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ (∈-++⁻ xs p) ≡ p - ∈-++⁻∘++⁺ : ∀ {v} xs {ys} (p : v ∈ xs ⊎ v ∈ ys) → ∈-++⁻ xs ([ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ p) ≡ p - ∈-++↔ : ∀ {v xs ys} → (v ∈ xs ⊎ v ∈ ys) ↔ v ∈ xs ++ ys - ∈-++-comm : ∀ {v} xs ys → v ∈ xs ++ ys → v ∈ ys ++ xs - ∈-++-comm∘++-comm : ∀ {v} xs {ys} (p : v ∈ xs ++ ys) → ∈-++-comm ys xs (∈-++-comm xs ys p) ≡ p - ∈-++↔++ : ∀ {v} xs ys → v ∈ xs ++ ys ↔ v ∈ ys ++ xs - ``` - -* Exposed container combinator conversion functions from `Data.Container.Combinator.Properties` separate from their correctness proofs in `Data.Container.Combinator`: - ``` - to-id : F.id A → ⟦ id ⟧ A - from-id : ⟦ id ⟧ A → F.id A - to-const : (A : Set s) → A → ⟦ const A ⟧ B - from-const : (A : Set s) → ⟦ const A ⟧ B → A - to-∘ : (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) → ⟦ C₁ ⟧ (⟦ C₂ ⟧ A) → ⟦ C₁ ∘ C₂ ⟧ A - from-∘ : (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) → ⟦ C₁ ∘ C₂ ⟧ A → ⟦ C₁ ⟧ (⟦ C₂ ⟧ A) - to-× : (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) → ⟦ C₁ ⟧ A P.× ⟦ C₂ ⟧ A → ⟦ C₁ × C₂ ⟧ A - from-× : (C₁ : Container s₁ p₁) (C₂ : Container s₂ p₂) → ⟦ C₁ × C₂ ⟧ A → ⟦ C₁ ⟧ A P.× ⟦ C₂ ⟧ A - to-Π : (I : Set i) (Cᵢ : I → Container s p) → (∀ i → ⟦ Cᵢ i ⟧ A) → ⟦ Π I Cᵢ ⟧ A - from-Π : (I : Set i) (Cᵢ : I → Container s p) → ⟦ Π I Cᵢ ⟧ A → ∀ i → ⟦ Cᵢ i ⟧ A - to-⊎ : (C₁ : Container s₁ p) (C₂ : Container s₂ p) → ⟦ C₁ ⟧ A S.⊎ ⟦ C₂ ⟧ A → ⟦ C₁ ⊎ C₂ ⟧ A - from-⊎ : (C₁ : Container s₁ p) (C₂ : Container s₂ p) → ⟦ C₁ ⊎ C₂ ⟧ A → ⟦ C₁ ⟧ A S.⊎ ⟦ C₂ ⟧ A - to-Σ : (I : Set i) (C : I → Container s p) → (∃ λ i → ⟦ C i ⟧ A) → ⟦ Σ I C ⟧ A - from-Σ : (I : Set i) (C : I → Container s p) → ⟦ Σ I C ⟧ A → ∃ λ i → ⟦ C i ⟧ A - ``` - -* Added a non-dependent version of `Function.Base.flip` due to an issue noted in - Pull Request #1812: - ```agda - flip′ : (A → B → C) → (B → A → C) - ``` - -* Added new proofs to `Data.List.Properties` - ```agda - cartesianProductWith-zeroˡ : cartesianProductWith f [] ys ≡ [] - cartesianProductWith-zeroʳ : cartesianProductWith f xs [] ≡ [] - cartesianProductWith-distribʳ-++ : cartesianProductWith f (xs ++ ys) zs ≡ - cartesianProductWith f xs zs ++ cartesianProductWith f ys zs - foldr-map : foldr f x (map g xs) ≡ foldr (g -⟨ f ∣) x xs - foldl-map : foldl f x (map g xs) ≡ foldl (∣ f ⟩- g) x xs - ``` - -NonZero/Positive/Negative changes ---------------------------------- - -This is a full list of proofs that have changed form to use irrelevant instance arguments: - -* In `Data.Nat.Base`: - ``` - ≢-nonZero⁻¹ : ∀ {n} → .(NonZero n) → n ≢ 0 - >-nonZero⁻¹ : ∀ {n} → .(NonZero n) → n > 0 - ``` - -* In `Data.Nat.Properties`: - ``` - *-cancelʳ-≡ : ∀ m n {o} → m * suc o ≡ n * suc o → m ≡ n - *-cancelˡ-≡ : ∀ {m n} o → suc o * m ≡ suc o * n → m ≡ n - *-cancelʳ-≤ : ∀ m n o → m * suc o ≤ n * suc o → m ≤ n - *-cancelˡ-≤ : ∀ {m n} o → suc o * m ≤ suc o * n → m ≤ n - *-monoˡ-< : ∀ n → (_* suc n) Preserves _<_ ⟶ _<_ - *-monoʳ-< : ∀ n → (suc n *_) Preserves _<_ ⟶ _<_ - - m≤m*n : ∀ m {n} → 0 < n → m ≤ m * n - m≤n*m : ∀ m {n} → 0 < n → m ≤ n * m - m⇒∤ : ∀ {m n} → m > suc n → m ∤ suc n - *-cancelˡ-∣ : ∀ {i j} k → suc k * i ∣ suc k * j → i ∣ j - ``` - -* In `Data.Nat.DivMod`: - ``` - m≡m%n+[m/n]*n : ∀ m n → m ≡ m % suc n + (m / suc n) * suc n - m%n≡m∸m/n*n : ∀ m n → m % suc n ≡ m ∸ (m / suc n) * suc n - n%n≡0 : ∀ n → suc n % suc n ≡ 0 - m%n%n≡m%n : ∀ m n → m % suc n % suc n ≡ m % suc n - [m+n]%n≡m%n : ∀ m n → (m + suc n) % suc n ≡ m % suc n - [m+kn]%n≡m%n : ∀ m k n → (m + k * (suc n)) % suc n ≡ m % suc n - m*n%n≡0 : ∀ m n → (m * suc n) % suc n ≡ 0 - m%n 0ℤ - negative⁻¹ : ∀ {i} → Negative i → i < 0ℤ - nonPositive⁻¹ : ∀ {i} → NonPositive i → i ≤ 0ℤ - nonNegative⁻¹ : ∀ {i} → NonNegative i → i ≥ 0ℤ - negative_ - *-monoʳ-<-neg : ∀ n → (_* -[1+ n ]) Preserves _<_ ⟶ _>_ - *-cancelˡ-<-nonPos : ∀ n → NonPositive n → n * i < n * j → i > j - *-cancelʳ-<-nonPos : ∀ n → NonPositive n → i * n < j * n → i > j - - *-distribˡ-⊓-nonNeg : ∀ m j k → + m * (j ⊓ k) ≡ (+ m * j) ⊓ (+ m * k) - *-distribʳ-⊓-nonNeg : ∀ m j k → (j ⊓ k) * + m ≡ (j * + m) ⊓ (k * + m) - *-distribˡ-⊓-nonPos : ∀ i → NonPositive i → ∀ j k → i * (j ⊓ k) ≡ (i * j) ⊔ (i * k) - *-distribʳ-⊓-nonPos : ∀ i → NonPositive i → ∀ j k → (j ⊓ k) * i ≡ (j * i) ⊔ (k * i) - *-distribˡ-⊔-nonNeg : ∀ m j k → + m * (j ⊔ k) ≡ (+ m * j) ⊔ (+ m * k) - *-distribʳ-⊔-nonNeg : ∀ m j k → (j ⊔ k) * + m ≡ (j * + m) ⊔ (k * + m) - *-distribˡ-⊔-nonPos : ∀ i → NonPositive i → ∀ j k → i * (j ⊔ k) ≡ (i * j) ⊓ (i * k) - *-distribʳ-⊔-nonPos : ∀ i → NonPositive i → ∀ j k → (j ⊔ k) * i ≡ (j * i) ⊓ (k * i) - ``` - -* In `Data.Integer.Divisibility`: - ``` - *-cancelˡ-∣ : ∀ k {i j} → k ≢ + 0 → k * i ∣ k * j → i ∣ j - *-cancelʳ-∣ : ∀ k {i j} → k ≢ + 0 → i * k ∣ j * k → i ∣ j - ``` - -* In `Data.Integer.Divisibility.Signed`: - ``` - *-cancelˡ-∣ : ∀ k {i j} → k ≢ + 0 → k * i ∣ k * j → i ∣ j - *-cancelʳ-∣ : ∀ k {i j} → k ≢ + 0 → i * k ∣ j * k → i ∣ j - ``` - -* In `Data.Rational.Unnormalised.Properties`: - ```agda - positive⁻¹ : ∀ {q} → .(Positive q) → q > 0ℚᵘ - nonNegative⁻¹ : ∀ {q} → .(NonNegative q) → q ≥ 0ℚᵘ - negative⁻¹ : ∀ {q} → .(Negative q) → q < 0ℚᵘ - nonPositive⁻¹ : ∀ {q} → .(NonPositive q) → q ≤ 0ℚᵘ - positive⇒nonNegative : ∀ {p} → Positive p → NonNegative p - negative⇒nonPositive : ∀ {p} → Negative p → NonPositive p - negative_ - *-monoʳ-<-neg : ∀ r → Negative r → (r *_) Preserves _<_ ⟶ _>_ - - pos⇒1/pos : ∀ p (p>0 : Positive p) → Positive ((1/ p) {{pos⇒≢0 p p>0}}) - neg⇒1/neg : ∀ p (p<0 : Negative p) → Negative ((1/ p) {{neg⇒≢0 p p<0}}) - - *-distribʳ-⊓-nonNeg : ∀ p → NonNegative p → ∀ q r → (q ⊓ r) * p ≃ (q * p) ⊓ (r * p) - *-distribˡ-⊓-nonNeg : ∀ p → NonNegative p → ∀ q r → p * (q ⊓ r) ≃ (p * q) ⊓ (p * r) - *-distribˡ-⊔-nonNeg : ∀ p → NonNegative p → ∀ q r → p * (q ⊔ r) ≃ (p * q) ⊔ (p * r) - *-distribʳ-⊔-nonNeg : ∀ p → NonNegative p → ∀ q r → (q ⊔ r) * p ≃ (q * p) ⊔ (r * p) - *-distribˡ-⊔-nonPos : ∀ p → NonPositive p → ∀ q r → p * (q ⊔ r) ≃ (p * q) ⊓ (p * r) - *-distribʳ-⊔-nonPos : ∀ p → NonPositive p → ∀ q r → (q ⊔ r) * p ≃ (q * p) ⊓ (r * p) - *-distribˡ-⊓-nonPos : ∀ p → NonPositive p → ∀ q r → p * (q ⊓ r) ≃ (p * q) ⊔ (p * r) - *-distribʳ-⊓-nonPos : ∀ p → NonPositive p → ∀ q r → (q ⊓ r) * p ≃ (q * p) ⊔ (r * p) - ``` - -* In `Data.Rational.Properties`: - ``` - positive⁻¹ : Positive p → p > 0ℚ - nonNegative⁻¹ : NonNegative p → p ≥ 0ℚ - negative⁻¹ : Negative p → p < 0ℚ - nonPositive⁻¹ : NonPositive p → p ≤ 0ℚ - negative q - *-cancelʳ-<-nonPos : ∀ r → NonPositive r → ∀ {p q} → p * r < q * r → p > q - *-monoʳ-≤-nonNeg : ∀ r → NonNegative r → (_* r) Preserves _≤_ ⟶ _≤_ - *-monoˡ-≤-nonNeg : ∀ r → NonNegative r → (r *_) Preserves _≤_ ⟶ _≤_ - *-monoʳ-≤-nonPos : ∀ r → NonPositive r → (_* r) Preserves _≤_ ⟶ _≥_ - *-monoˡ-≤-nonPos : ∀ r → NonPositive r → (r *_) Preserves _≤_ ⟶ _≥_ - *-monoˡ-<-pos : ∀ r → Positive r → (_* r) Preserves _<_ ⟶ _<_ - *-monoʳ-<-pos : ∀ r → Positive r → (r *_) Preserves _<_ ⟶ _<_ - *-monoˡ-<-neg : ∀ r → Negative r → (_* r) Preserves _<_ ⟶ _>_ - *-monoʳ-<-neg : ∀ r → Negative r → (r *_) Preserves _<_ ⟶ _>_ - - *-distribˡ-⊓-nonNeg : ∀ p → NonNegative p → ∀ q r → p * (q ⊓ r) ≡ (p * q) ⊓ (p * r) - *-distribʳ-⊓-nonNeg : ∀ p → NonNegative p → ∀ q r → (q ⊓ r) * p ≡ (q * p) ⊓ (r * p) - *-distribˡ-⊔-nonNeg : ∀ p → NonNegative p → ∀ q r → p * (q ⊔ r) ≡ (p * q) ⊔ (p * r) - *-distribʳ-⊔-nonNeg : ∀ p → NonNegative p → ∀ q r → (q ⊔ r) * p ≡ (q * p) ⊔ (r * p) - *-distribˡ-⊔-nonPos : ∀ p → NonPositive p → ∀ q r → p * (q ⊔ r) ≡ (p * q) ⊓ (p * r) - *-distribʳ-⊔-nonPos : ∀ p → NonPositive p → ∀ q r → (q ⊔ r) * p ≡ (q * p) ⊓ (r * p) - *-distribˡ-⊓-nonPos : ∀ p → NonPositive p → ∀ q r → p * (q ⊓ r) ≡ (p * q) ⊔ (p * r) - *-distribʳ-⊓-nonPos : ∀ p → NonPositive p → ∀ q r → (q ⊓ r) * p ≡ (q * p) ⊔ (r * p) - - pos⇒1/pos : ∀ p (p>0 : Positive p) → Positive ((1/ p) {{pos⇒≢0 p p>0}}) - neg⇒1/neg : ∀ p (p<0 : Negative p) → Negative ((1/ p) {{neg⇒≢0 p p<0}}) - 1/pos⇒pos : ∀ p .{{_ : NonZero p}} → (1/p : Positive (1/ p)) → Positive p - 1/neg⇒neg : ∀ p .{{_ : NonZero p}} → (1/p : Negative (1/ p)) → Negative p - ``` - -* In `Data.List.NonEmpty.Base`: - ```agda - drop+ : ℕ → List⁺ A → List⁺ A - ``` - When drop+ping more than the size of the length of the list, the last element remains. - -* Added new proofs in `Data.List.NonEmpty.Properties`: - ```agda - length-++⁺ : length (xs ++⁺ ys) ≡ length xs + length ys - length-++⁺-tail : length (xs ++⁺ ys) ≡ suc (length xs + length (tail ys)) - ++-++⁺ : (xs ++ ys) ++⁺ zs ≡ xs ++⁺ ys ++⁺ zs - ++⁺-cancelˡ′ : xs ++⁺ zs ≡ ys ++⁺ zs′ → List.length xs ≡ List.length ys → zs ≡ zs′ - ++⁺-cancelˡ : xs ++⁺ ys ≡ xs ++⁺ zs → ys ≡ zs - drop+-++⁺ : drop+ (length xs) (xs ++⁺ ys) ≡ ys - map-++⁺-commute : map f (xs ++⁺ ys) ≡ map f xs ++⁺ map f ys - length-map : length (map f xs) ≡ length xs - map-cong : f ≗ g → map f ≗ map g - map-compose : map (g ∘ f) ≗ map g ∘ map f - ``` - -* Added new functions and proofs in `Data.Nat.GeneralisedArithmetic`: - ```agda - iterate : (A → A) → A → ℕ → A - iterate-is-fold : ∀ (z : A) s m → fold z s m ≡ iterate s z m - ``` - -* Added new proofs to `Function.Properties.Inverse`: - ```agda - Inverse⇒Injection : Inverse S T → Injection S T - ↔⇒↣ : A ↔ B → A ↣ B - ``` - -* Added a new isomorphism to `Data.Fin.Properties`: - ```agda - 2↔Bool : Fin 2 ↔ Bool - ``` - -* Added new isomorphisms to `Data.Unit.Polymorphic.Properties`: - ```agda - ⊤↔⊤* : ⊤ {ℓ} ↔ ⊤* - ``` - -* Added new isomorphisms to `Data.Vec.N-ary`: - ```agda - Vec↔N-ary : ∀ n → (Vec A n → B) ↔ N-ary n A B - ``` - -* Added new isomorphisms to `Data.Vec.Recursive`: - ```agda - lift↔ : ∀ n → A ↔ B → A ^ n ↔ B ^ n - Fin[m^n]↔Fin[m]^n : ∀ m n → Fin (m ^ n) ↔ Fin m Vec.^ n - ``` - -* Added new functions to `Function.Properties.Inverse`: - ```agda - ↔-refl : Reflexive _↔_ - ↔-sym : Symmetric _↔_ - ↔-trans : Transitive _↔_ - ``` - -* Added new isomorphisms to `Function.Properties.Inverse`: - ```agda - ↔-fun : A ↔ B → C ↔ D → (A → C) ↔ (B → D) - ``` - -* Added new function to `Data.Fin.Properties` - ```agda - i≤inject₁[j]⇒i≤1+j : i ≤ inject₁ j → i ≤ suc j - ``` - -* Added new function to `Data.Fin.Induction` - ```agda - <-weakInduction-startingFrom : P i → (∀ j → P (inject₁ j) → P (suc j)) → ∀ {j} → j ≥ i → P j - ``` - -* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`: - ```agda - foldr-commMonoid : xs ↭ ys → foldr _∙_ ε xs ≈ foldr _∙_ ε ys - ``` - -* Added new module to `Data.Rational.Unnormalised.Properties` - ```agda - module ≃-Reasoning = SetoidReasoning ≃-setoid - ``` - -* Added new functions to `Data.Rational.Unnormalised.Properties` - ```agda - 0≠1 : 0ℚᵘ ≠ 1ℚᵘ - ≃-≠-irreflexive : Irreflexive _≃_ _≠_ - ≠-symmetric : Symmetric _≠_ - ≠-cotransitive : Cotransitive _≠_ - ≠⇒invertible : p ≠ q → Invertible _≃_ 1ℚᵘ _*_ (p - q) - ``` - -* Added new structures to `Data.Rational.Unnormalised.Properties` - ```agda - +-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ - +-*-isHeytingField : IsHeytingField _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ - ``` - -* Added new bundles to `Data.Rational.Unnormalised.Properties` - ```agda - +-*-heytingCommutativeRing : HeytingCommutativeRing 0ℓ 0ℓ 0ℓ - +-*-heytingField : HeytingField 0ℓ 0ℓ 0ℓ - ``` - -* Added new function to `Data.Vec.Relation.Binary.Pointwise.Inductive` - ```agda - cong-[_]≔ : Pointwise _∼_ xs ys → Pointwise _∼_ (xs [ i ]≔ p) (ys [ i ]≔ p) - ``` - -* Added new function to `Data.Vec.Relation.Binary.Equality.Setoid` - ```agda - map-[]≔ : map f (xs [ i ]≔ p) ≋ map f xs [ i ]≔ f p - ``` - -* Added new function to `Data.List.Relation.Binary.Permutation.Propositional.Properties` - ```agda - ↭-reverse : (xs : List A) → reverse xs ↭ xs - ``` - -* Added new functions to `Algebra.Properties.CommutativeMonoid` - ```agda - invertibleˡ⇒invertibleʳ : LeftInvertible _≈_ 0# _+_ x → RightInvertible _≈_ 0# _+_ x - invertibleʳ⇒invertibleˡ : RightInvertible _≈_ 0# _+_ x → LeftInvertible _≈_ 0# _+_ x - invertibleˡ⇒invertible : LeftInvertible _≈_ 0# _+_ x → Invertible _≈_ 0# _+_ x - invertibleʳ⇒invertible : RightInvertible _≈_ 0# _+_ x → Invertible _≈_ 0# _+_ x - ``` - -* Added new functions to `Algebra.Apartness.Bundles` - ```agda - invertibleˡ⇒# : LeftInvertible _≈_ 1# _*_ (x - y) → x # y - invertibleʳ⇒# : RightInvertible _≈_ 1# _*_ (x - y) → x # y - x#0y#0→xy#0 : x # 0# → y # 0# → x * y # 0# - #-sym : Symmetric _#_ - #-congʳ : x ≈ y → x # z → y # z - #-congˡ : y ≈ z → x # y → x # z - ``` - -* Added new proof to `Data.List.Relation.Unary.All.Properties`: - ```agda - gmap⁻ : Q ∘ f ⋐ P → All Q ∘ map f ⋐ All P - ``` - -* Added new proofs to `Data.List.Relation.Binary.Sublist.Setoid.Properties` - and `Data.List.Relation.Unary.Sorted.TotalOrder.Properties`. - ```agda - ⊆-mergeˡ : ∀ xs ys → xs ⊆ merge _≤?_ xs ys - ⊆-mergeʳ : ∀ xs ys → ys ⊆ merge _≤?_ xs ys - ``` - -* Added new proof to `Induction.WellFounded` - ```agda - Acc-resp-flip-≈ : _<_ Respectsʳ (flip _≈_) → (Acc _<_) Respects _≈_ - ``` - -* Added new file `Relation.Binary.Reasoning.Base.Apartness` - This is how to use it: - ```agda - _ : a # d - _ = begin-apartness - a ≈⟨ a≈b ⟩ - b #⟨ b#c ⟩ - c ≈⟨ c≈d ⟩ - d ∎ - ``` +* In `Function.Bundles`, added `_➙_` as a synonym for `Func` that can + be used infix. From 30e6ca57e198867f2dcddab75aeb9974cb6b1030 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 2 Jan 2024 17:49:10 +0000 Subject: [PATCH 25/27] added new defns and modules to `CHANGELOG` --- CHANGELOG.md | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index 8be73afc66..ee1f4a02d9 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -29,9 +29,31 @@ Deprecated names New modules ----------- +* Algebraic structures obtained as the free thing (for their signature): + ``` + Algebra.Free + Algebra.Free.Magma + ``` + +* Bundled morphisms between algebraic structures: + ``` + Algebra.Morphism.Bundles + ``` + Additions to existing modules ----------------------------- +* In `Algebra.Morphism.Construct.Composition`: + ``` + magmaHomomorphism : MagmaHomomorphism M₁ M₂ → MagmaHomomorphism M₂ M₃ → + MagmaHomomorphism M₁ M₃ + ``` + +* In `Algebra.Morphism.Construct.Identity`: + ``` + magmaHomomorphism : MagmaHomomorphism M M + ``` + * In `Data.Fin.Properties`: ```agda nonZeroIndex : Fin n → ℕ.NonZero n From 1dbc6fc66e709c2d7fb891967274aad23734f263 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 2 Jan 2024 18:01:01 +0000 Subject: [PATCH 26/27] use new symmetric equality reasoning combinator --- src/Algebra/Free/Magma.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index 8bfd2fca48..c742ff4786 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -326,7 +326,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) open Uniqueness 𝓚 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴷ) isUnique⟦_⟧ : ∀ t → ⟦ t ⟧ᴴ M.≈ ⟦ t ⟧ᴷ - isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧η ≈˘⟨ isUnique⟦ t ⟧ᴷ ⟩ ⟦ t ⟧ᴷ ∎ + isUnique⟦ t ⟧ = begin ⟦ t ⟧ᴴ ≈⟨ isUnique⟦ t ⟧ᴴ ⟩ ⟦ t ⟧η ≈⟨ isUnique⟦ t ⟧ᴷ ⟨ ⟦ t ⟧ᴷ ∎ ------------------------------------------------------------------------ -- Immediate corollary: alg is in fact a MagmaHomomorphism From 9782cbf14640f8f373dd529643ac1dcc09a7d58e Mon Sep 17 00:00:00 2001 From: James McKinna Date: Mon, 8 Apr 2024 08:40:09 +0100 Subject: [PATCH 27/27] `fix-whitespace` --- src/Algebra/Free/Magma.agda | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/Algebra/Free/Magma.agda b/src/Algebra/Free/Magma.agda index c742ff4786..324fd380b9 100644 --- a/src/Algebra/Free/Magma.agda +++ b/src/Algebra/Free/Magma.agda @@ -46,7 +46,7 @@ private 𝓒 : Setoid c ℓc 𝓜 : Magma m ℓm 𝓝 : Magma n ℓn - + ------------------------------------------------------------------------ -- Syntax: 'pre'-free algebra @@ -319,7 +319,7 @@ module LeftAdjoint {𝓐 : Setoid a ℓa} (𝓜 : Magma m ℓm) ⟦ s ⟧η M.∙ ⟦ t ⟧η ∎ module Corollary (𝓗 𝓚 : η-MagmaHomomorphism) where - + open η-MagmaHomomorphism 𝓗 using () renaming (⟦_⟧ to ⟦_⟧ᴴ) open η-MagmaHomomorphism 𝓚 using () renaming (⟦_⟧ to ⟦_⟧ᴷ) open Uniqueness 𝓗 renaming (isUnique⟦_⟧ to isUnique⟦_⟧ᴴ) @@ -348,7 +348,7 @@ module FreeMagmaFunctor (𝓗 : SetoidHomomorphism 𝓐 𝓑) where private module FA = FreeMagma 𝓐 module FB = FreeMagma 𝓑 - + magmaHomomorphism : MagmaHomomorphism FA.magma FB.magma magmaHomomorphism = Existence.magmaHomomorphism where open LeftAdjoint FB.magma (Compose.setoidHomomorphism 𝓗 FB.varSetoidHomomorphism) @@ -358,14 +358,14 @@ module FreeMagmaFunctor (𝓗 : SetoidHomomorphism 𝓐 𝓑) where module Naturality (𝓕 : MagmaHomomorphism 𝓜 𝓝) where - private + private module M = Magma 𝓜 module N = Magma 𝓝 module F = MagmaHomomorphism 𝓕 Free𝓕 = FreeMagmaFunctor.magmaHomomorphism (F.setoidHomomorphism) module AlgM = MagmaHomomorphism (algMagmaHomomorphism 𝓜) module AlgN = MagmaHomomorphism (algMagmaHomomorphism 𝓝) - + module Map = MagmaHomomorphism Free𝓕 naturality : ∀ t → F.⟦ AlgM.⟦ t ⟧ ⟧ N.≈ AlgN.⟦ Map.⟦ t ⟧ ⟧