From 63d6ca0cc9a1d63b1d5d93264790580ba285c982 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 7 Oct 2023 06:41:54 +0100 Subject: [PATCH 01/57] basic result --- src/Induction/NoInfiniteDescent.agda | 56 ++++++++++++++++++++++++++++ 1 file changed, 56 insertions(+) create mode 100644 src/Induction/NoInfiniteDescent.agda diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda new file mode 100644 index 0000000000..c10f102f24 --- /dev/null +++ b/src/Induction/NoInfiniteDescent.agda @@ -0,0 +1,56 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- A standard consequence of accessibility/well-foundedness: +-- the impossibility of 'infinite descent' from any x st P to +-- 'smaller' y also satisfying P +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Induction.NoInfiniteDescent where + +open import Data.Product.Base using (_,_; proj₁; proj₂; ∃-syntax; _×_) +open import Function.Base using (_∘_) +open import Induction.WellFounded +open import Level using (Level; _⊔_) +open import Relation.Binary.Core using (Rel) +open import Relation.Binary.Construct.Closure.Transitive +open import Relation.Nullary.Negation.Core using (¬_) +open import Relation.Unary using (Pred) + +private + variable + a b ℓ ℓ₁ ℓ₂ r : Level + A : Set a + B : Set b + +------------------------------------------------------------------------ +-- Definitions + +module _ {_<_ : Rel A r} (P : Pred A ℓ) where + + private + _<⁺_ : Rel A _ + z <⁺ x = z ⟨ _<_ ⟩⁺ x + + SmallerCounterexample : Pred A _ + SmallerCounterexample x = (px : P x) → ∃[ y ] y < x × P y + + SmallerCounterexample⁺ : Pred A _ + SmallerCounterexample⁺ x = (px : P x) → ∃[ y ] y <⁺ x × P y + +------------------------------------------------------------------------ +-- Basic result + + module Lemma (sce : ∀ {x} → SmallerCounterexample x) where + + accNoSmallestCounterexample : ∀ {x} → Acc _<_ x → ¬ (P x) + accNoSmallestCounterexample = Some.wfRec _ + (λ _ hy py → let z , z Date: Sat, 7 Oct 2023 06:53:16 +0100 Subject: [PATCH 02/57] added missing lemma; is there a better name for this? --- src/Relation/Binary/Construct/Closure/Transitive.agda | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/Relation/Binary/Construct/Closure/Transitive.agda b/src/Relation/Binary/Construct/Closure/Transitive.agda index e4154d49cd..5ece75b885 100644 --- a/src/Relation/Binary/Construct/Closure/Transitive.agda +++ b/src/Relation/Binary/Construct/Closure/Transitive.agda @@ -70,6 +70,13 @@ module _ (_∼_ : Rel A ℓ) where transitive : Transitive _∼⁺_ transitive = _++_ + transitive⇒∼⁺⊆∼ : Transitive _∼_ → _∼⁺_ ⇒ _∼_ + transitive⇒∼⁺⊆∼ trans = ∼⁺⊆∼ + where + ∼⁺⊆∼ : _∼⁺_ ⇒ _∼_ + ∼⁺⊆∼ [ x∼y ] = x∼y + ∼⁺⊆∼ (x∼y ∷ x∼⁺y) = trans x∼y (∼⁺⊆∼ x∼⁺y) + accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x accessible⁻ = ∼⊆∼⁺.accessible From 60d22e650132bab4b38eef4eff230a3766b961dc Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 7 Oct 2023 06:56:39 +0100 Subject: [PATCH 03/57] renamed lemma --- src/Relation/Binary/Construct/Closure/Transitive.agda | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/Relation/Binary/Construct/Closure/Transitive.agda b/src/Relation/Binary/Construct/Closure/Transitive.agda index 5ece75b885..9741d921a0 100644 --- a/src/Relation/Binary/Construct/Closure/Transitive.agda +++ b/src/Relation/Binary/Construct/Closure/Transitive.agda @@ -70,11 +70,11 @@ module _ (_∼_ : Rel A ℓ) where transitive : Transitive _∼⁺_ transitive = _++_ - transitive⇒∼⁺⊆∼ : Transitive _∼_ → _∼⁺_ ⇒ _∼_ - transitive⇒∼⁺⊆∼ trans = ∼⁺⊆∼ + transitive⁻ : Transitive _∼_ → _∼⁺_ ⇒ _∼_ + transitive⁻ trans = ∼⁺⊆∼ where ∼⁺⊆∼ : _∼⁺_ ⇒ _∼_ - ∼⁺⊆∼ [ x∼y ] = x∼y + ∼⁺⊆∼ [ x∼y ] = x∼y ∼⁺⊆∼ (x∼y ∷ x∼⁺y) = trans x∼y (∼⁺⊆∼ x∼⁺y) accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x From ec0966940c5111d1318347bdc2db5dc3523f2af8 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 7 Oct 2023 10:48:05 +0100 Subject: [PATCH 04/57] final tidying up; `CHANGELOG` --- CHANGELOG.md | 5 ++ src/Induction/NoInfiniteDescent.agda | 73 +++++++++++++++++++--------- 2 files changed, 54 insertions(+), 24 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 3d99a6eb2b..f72763f1ba 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1904,6 +1904,11 @@ New modules Function.Properties.Surjection ``` +* The 'no infinite descent' principle for (accessible elements of) well-founded relations: + ``` + Induction.NoInfiniteDescent + ``` + * In order to improve modularity, the contents of `Relation.Binary.Lattice` has been split out into the standard: ``` diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index c10f102f24..61cef555ad 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -2,55 +2,80 @@ -- The Agda standard library -- -- A standard consequence of accessibility/well-foundedness: --- the impossibility of 'infinite descent' from any x st P to --- 'smaller' y also satisfying P +-- the impossibility of 'infinite descent' from any (accessible) +-- element x satisfying P to 'smaller' y also satisfying P ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} module Induction.NoInfiniteDescent where -open import Data.Product.Base using (_,_; proj₁; proj₂; ∃-syntax; _×_) +open import Data.Product.Base using (_,_; ∃-syntax; _×_) open import Function.Base using (_∘_) -open import Induction.WellFounded -open import Level using (Level; _⊔_) +open import Induction.WellFounded using (WellFounded; Acc; acc-inverse; module Some) +open import Level using (Level) open import Relation.Binary.Core using (Rel) -open import Relation.Binary.Construct.Closure.Transitive open import Relation.Nullary.Negation.Core using (¬_) -open import Relation.Unary using (Pred) +open import Relation.Unary using (Pred; _∩_) private variable - a b ℓ ℓ₁ ℓ₂ r : Level + a r ℓ : Level A : Set a - B : Set b ------------------------------------------------------------------------ -- Definitions -module _ {_<_ : Rel A r} (P : Pred A ℓ) where +module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where - private - _<⁺_ : Rel A _ - z <⁺ x = z ⟨ _<_ ⟩⁺ x + DescentAt : Pred A _ + DescentAt x = P x → ∃[ y ] y < x × P y - SmallerCounterexample : Pred A _ - SmallerCounterexample x = (px : P x) → ∃[ y ] y < x × P y + InfiniteDescent : Set _ + InfiniteDescent = ∀ {x} → DescentAt x - SmallerCounterexample⁺ : Pred A _ - SmallerCounterexample⁺ x = (px : P x) → ∃[ y ] y <⁺ x × P y + AccInfiniteDescent : Set _ + AccInfiniteDescent = ∀ {x} → Acc _<_ x → DescentAt x ------------------------------------------------------------------------ --- Basic result +-- Basic result: assumes unrestricted descent - module Lemma (sce : ∀ {x} → SmallerCounterexample x) where + module Lemma (descent : InfiniteDescent) where - accNoSmallestCounterexample : ∀ {x} → Acc _<_ x → ¬ (P x) - accNoSmallestCounterexample = Some.wfRec _ - (λ _ hy py → let z , z Date: Sat, 7 Oct 2023 10:50:19 +0100 Subject: [PATCH 05/57] final tidying up --- CHANGELOG.md | 20 ++++++++++---------- src/Induction/NoInfiniteDescent.agda | 2 +- 2 files changed, 11 insertions(+), 11 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index f72763f1ba..99fd1f233f 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -599,7 +599,7 @@ Non-backwards compatible changes with the consequence that all arguments involving about accesibility and wellfoundedness proofs were polluted by almost-always-inferrable explicit arguments for the `y` position. The definition has now been changed to - make that argument *implicit*, as + make that argument *implicit*, as ```agda WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _ WfRec _<_ P x = ∀ {y} → y < x → P y @@ -621,7 +621,7 @@ Non-backwards compatible changes raw bundles to `Data.X.Base`, this definition can now be made directly. Knock-on consequences include the need to retain the old constructor name, now introduced as a pattern synonym, and introduction of (a function equivalent to) the former - field name/projection function `proof` as `≤″-proof` in `Data.Nat.Properties`. + field name/projection function `proof` as `≤″-proof` in `Data.Nat.Properties`. * Accordingly, the definition has been changed to: ```agda @@ -879,12 +879,12 @@ Non-backwards compatible changes ``` _─_ ↦ removeAt ``` - + * In `Data.Vec.Base`: ```agda insert ↦ insertAt remove ↦ removeAt - + updateAt : Fin n → (A → A) → Vec A n → Vec A n ↦ updateAt : Vec A n → Fin n → (A → A) → Vec A n @@ -895,12 +895,12 @@ Non-backwards compatible changes remove : Fin (suc n) → Vector A (suc n) → Vector A n ↦ removeAt : Vector A (suc n) → Fin (suc n) → Vector A n - + updateAt : Fin n → (A → A) → Vector A n → Vector A n ↦ updateAt : Vector A n → Fin n → (A → A) → Vector A n ``` - + * The old names (and the names of all proofs about these functions) have been deprecated appropriately. ### Changes to triple reasoning interface @@ -925,7 +925,7 @@ Non-backwards compatible changes Data.Vec.Relation.Binary.Lex.NonStrict Relation.Binary.Reasoning.StrictPartialOrder Relation.Binary.Reasoning.PartialOrder - ``` + ``` ### Other @@ -1103,7 +1103,7 @@ Non-backwards compatible changes * `excluded-middle` in `Relation.Nullary.Decidable.Core` has been renamed to `¬¬-excluded-middle`. -* `iterate` and `replicate` in `Data.Vec.Base` and `Data.Vec.Functional` +* `iterate` and `replicate` in `Data.Vec.Base` and `Data.Vec.Functional` now take the length of vector, `n`, as an explicit rather than an implicit argument. ```agda iterate : (A → A) → A → ∀ n → Vec A n @@ -1658,7 +1658,7 @@ Deprecated names take-drop-id ↦ take++drop≡id map-insert ↦ map-insertAt - + insert-lookup ↦ insertAt-lookup insert-punchIn ↦ insertAt-punchIn remove-PunchOut ↦ removeAt-punchOut @@ -1684,7 +1684,7 @@ Deprecated names updateAt-cong-relative ↦ updateAt-cong-local map-updateAt ↦ map-updateAt-local - + insert-lookup ↦ insertAt-lookup insert-punchIn ↦ insertAt-punchIn remove-punchOut ↦ removeAt-punchOut diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 61cef555ad..ceeba8e4de 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -60,7 +60,7 @@ module _ {_<_ : Rel A r} (P : Pred A ℓ) where accNoInfiniteDescent : ∀ {x} → Acc _<_ x → ¬ (P x) accNoInfiniteDescent ax px = ID∩.Lemma.accNoInfiniteDescent descent∩ ax (px , ax) - + where P∩Acc : Pred A _ P∩Acc = P ∩ (Acc _<_) From 00e00837f38e2d18ded2c050251018234d044a79 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 7 Oct 2023 11:06:22 +0100 Subject: [PATCH 06/57] missing `CHANGELOG` entry --- CHANGELOG.md | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index 99fd1f233f..20829c7b93 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -3167,6 +3167,11 @@ Additions to existing modules _⋆ : _⟶₁_ ⇒ SymClosure _⟶₂_ → SymClosure _⟶₁_ ⇒ SymClosure _⟶₂_ ``` +* Added new proof in `Relation.Binary.Construct.Closure.Transitive`: + ``` + transitive⁻ : Transitive _∼_ → TransClosure _∼_ ⇒ _∼_ + ``` + * Added new proofs to `Relation.Binary.Lattice.Properties.{Join,Meet}Semilattice`: ```agda isPosemigroup : IsPosemigroup _≈_ _≤_ _∨_ From c1b3ed967f9dbca0a1cade8c8ce5340d869b05ab Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 8 Oct 2023 19:35:14 +0100 Subject: [PATCH 07/57] `InfiniteDescent` definition and proof --- src/Induction/NoInfiniteDescent.agda | 81 +++++++++++++++++++++------- 1 file changed, 61 insertions(+), 20 deletions(-) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index ceeba8e4de..3f83a8ce9f 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -10,13 +10,15 @@ module Induction.NoInfiniteDescent where -open import Data.Product.Base using (_,_; ∃-syntax; _×_) +open import Data.Nat.Base using (ℕ; zero; suc) +open import Data.Product.Base using (_,_; Σ-syntax; ∃-syntax; _×_) open import Function.Base using (_∘_) -open import Induction.WellFounded using (WellFounded; Acc; acc-inverse; module Some) +open import Induction.WellFounded using (WellFounded; Acc; acc; acc-inverse; module Some) open import Level using (Level) open import Relation.Binary.Core using (Rel) +open import Relation.Binary.PropositionalEquality.Core open import Relation.Nullary.Negation.Core using (¬_) -open import Relation.Unary using (Pred; _∩_) +open import Relation.Unary using (Pred; _∩_; _⇒_; IUniversal) private variable @@ -31,22 +33,53 @@ module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where DescentAt : Pred A _ DescentAt x = P x → ∃[ y ] y < x × P y - InfiniteDescent : Set _ - InfiniteDescent = ∀ {x} → DescentAt x + AccDescent : Set _ + AccDescent = ∀[ Acc _<_ ⇒ DescentAt ] + + Descent : Set _ + Descent = ∀[ DescentAt ] + + InfiniteDescentAt : Pred A _ + InfiniteDescentAt x = P x → Σ[ f ∈ (ℕ → A) ] f zero ≡ x × ∀ n → f (suc n) < f n AccInfiniteDescent : Set _ - AccInfiniteDescent = ∀ {x} → Acc _<_ x → DescentAt x + AccInfiniteDescent = ∀[ Acc _<_ ⇒ InfiniteDescentAt ] + + InfiniteDescent : Set _ + InfiniteDescent = ∀[ InfiniteDescentAt ] ------------------------------------------------------------------------ -- Basic result: assumes unrestricted descent - module Lemma (descent : InfiniteDescent) where + module Lemmas (descent : Descent) where - accNoInfiniteDescent : ∀ {x} → Acc _<_ x → ¬ (P x) - accNoInfiniteDescent = Some.wfRec (¬_ ∘ P) - (λ _ hy py → let z , z Date: Mon, 9 Oct 2023 00:14:27 +0100 Subject: [PATCH 08/57] revised `InfiniteDescent` definition and proof --- src/Induction/NoInfiniteDescent.agda | 16 ++++++++++------ 1 file changed, 10 insertions(+), 6 deletions(-) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 3f83a8ce9f..9501f6c166 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -40,7 +40,8 @@ module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where Descent = ∀[ DescentAt ] InfiniteDescentAt : Pred A _ - InfiniteDescentAt x = P x → Σ[ f ∈ (ℕ → A) ] f zero ≡ x × ∀ n → f (suc n) < f n + InfiniteDescentAt x = P x → + Σ[ f ∈ (ℕ → A) ] f zero ≡ x × ∀ n → f (suc n) < f n × P (f n) AccInfiniteDescent : Set _ AccInfiniteDescent = ∀[ Acc _<_ ⇒ InfiniteDescentAt ] @@ -59,16 +60,17 @@ module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where rec : _ rec y rec[y] py with z , z Date: Mon, 9 Oct 2023 11:31:34 +0100 Subject: [PATCH 09/57] major revision: renames things, plus additional corollaries --- src/Induction/NoInfiniteDescent.agda | 147 ++++++++++++++++++++------- 1 file changed, 111 insertions(+), 36 deletions(-) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 9501f6c166..0f33853514 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -10,15 +10,17 @@ module Induction.NoInfiniteDescent where -open import Data.Nat.Base using (ℕ; zero; suc) -open import Data.Product.Base using (_,_; Σ-syntax; ∃-syntax; _×_) +open import Data.Nat.Base as ℕ using (ℕ; zero; suc) +open import Data.Nat.Properties as ℕ +open import Data.Product.Base using (_,_; proj₁; ∃-syntax; _×_) open import Function.Base using (_∘_) open import Induction.WellFounded using (WellFounded; Acc; acc; acc-inverse; module Some) open import Level using (Level) open import Relation.Binary.Core using (Rel) +open import Relation.Binary.Construct.Closure.Transitive open import Relation.Binary.PropositionalEquality.Core open import Relation.Nullary.Negation.Core using (¬_) -open import Relation.Unary using (Pred; _∩_; _⇒_; IUniversal) +open import Relation.Unary using (Pred; _∩_; _⇒_; Universal; IUniversal) private variable @@ -33,65 +35,73 @@ module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where DescentAt : Pred A _ DescentAt x = P x → ∃[ y ] y < x × P y - AccDescent : Set _ - AccDescent = ∀[ Acc _<_ ⇒ DescentAt ] + Acc⇒Descent : Set _ + Acc⇒Descent = ∀[ Acc _<_ ⇒ DescentAt ] Descent : Set _ Descent = ∀[ DescentAt ] + InfiniteDescendingSequence : (f : ℕ → A) → Set _ + InfiniteDescendingSequence f = ∀ n → f (suc n) < f n + + InfiniteSequence_From_ : (f : ℕ → A) → Pred A _ + InfiniteSequence f From x = f zero ≡ x × InfiniteDescendingSequence f + InfiniteDescentAt : Pred A _ - InfiniteDescentAt x = P x → - Σ[ f ∈ (ℕ → A) ] f zero ≡ x × ∀ n → f (suc n) < f n × P (f n) + InfiniteDescentAt x = P x → ∃[ f ] Π[ P ∘ f ] × InfiniteSequence f From x - AccInfiniteDescent : Set _ - AccInfiniteDescent = ∀[ Acc _<_ ⇒ InfiniteDescentAt ] + Acc⇒InfiniteDescent : Set _ + Acc⇒InfiniteDescent = ∀[ Acc _<_ ⇒ InfiniteDescentAt ] InfiniteDescent : Set _ InfiniteDescent = ∀[ InfiniteDescentAt ] ------------------------------------------------------------------------ --- Basic result: assumes unrestricted descent +-- Basic lemmas: assume unrestricted descent module Lemmas (descent : Descent) where - accInfiniteDescent : AccInfiniteDescent - accInfiniteDescent {x} = Some.wfRec InfiniteDescentAt rec x + acc⇒infiniteDescent : Acc⇒InfiniteDescent + acc⇒infiniteDescent {x} = Some.wfRec InfiniteDescentAt rec x where rec : _ rec y rec[y] py with z , z Date: Mon, 9 Oct 2023 11:39:45 +0100 Subject: [PATCH 10/57] spacing --- src/Induction/NoInfiniteDescent.agda | 1 + 1 file changed, 1 insertion(+) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 0f33853514..cf5b398ac9 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -169,6 +169,7 @@ module _ {_<_ : Rel A r} (P : Pred A ℓ) where private descent : ID⁺.Acc⇒Descent descent = descent⁺ ∘ (accessible⁻ _) + module Corollaries⁺ = Corollaries _ descent sequence⁺ : ∀ {f} → From 17cc4b44d0c03817c74585ff5f241d768aeea3ce Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 15 Oct 2023 11:12:52 +0100 Subject: [PATCH 11/57] added `NoSmallestCounterExample` principles for `Stable` predicates --- src/Induction/NoInfiniteDescent.agda | 71 +++++++++++++++++++--------- 1 file changed, 48 insertions(+), 23 deletions(-) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index cf5b398ac9..066fb7013c 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -19,7 +19,7 @@ open import Level using (Level) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Construct.Closure.Transitive open import Relation.Binary.PropositionalEquality.Core -open import Relation.Nullary.Negation.Core using (¬_) +open import Relation.Nullary.Negation.Core as Negation using (¬_) open import Relation.Unary using (Pred; _∩_; _⇒_; Universal; IUniversal) private @@ -99,43 +99,43 @@ module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where module _ {_<_ : Rel A r} (P : Pred A ℓ) where - open InfiniteDescent _<_ P public + open InfiniteDescent _<_ P public - module Corollaries (descent : Acc⇒Descent) where + module Corollaries (descent : Acc⇒Descent) where - private + private - P∩Acc : Pred A _ - P∩Acc = P ∩ (Acc _<_) + P∩Acc : Pred A _ + P∩Acc = P ∩ (Acc _<_) - module ID∩ = InfiniteDescent _<_ P∩Acc + module ID∩ = InfiniteDescent _<_ P∩Acc - descent∩ : ID∩.Descent - descent∩ (px , acc[x]) = - let y , y Date: Sun, 15 Oct 2023 11:42:34 +0100 Subject: [PATCH 12/57] refactoring; move `Stable` to `Relation.Unary` --- CHANGELOG.md | 1 + src/Induction/NoInfiniteDescent.agda | 9 +++------ src/Relation/Unary.agda | 7 ++++++- 3 files changed, 10 insertions(+), 7 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 20829c7b93..4f65bf45c9 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -3213,6 +3213,7 @@ Additions to existing modules _≐_ : Pred A ℓ₁ → Pred A ℓ₂ → Set _ _≐′_ : Pred A ℓ₁ → Pred A ℓ₂ → Set _ _∖_ : Pred A ℓ₁ → Pred A ℓ₂ → Pred A _ + Stable : Pred A ℓ → Set _ ``` * Added new proofs in `Relation.Unary.Properties`: diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 066fb7013c..636e84c898 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -20,7 +20,7 @@ open import Relation.Binary.Core using (Rel) open import Relation.Binary.Construct.Closure.Transitive open import Relation.Binary.PropositionalEquality.Core open import Relation.Nullary.Negation.Core as Negation using (¬_) -open import Relation.Unary using (Pred; _∩_; _⇒_; Universal; IUniversal) +open import Relation.Unary using (Pred; _∩_; _⇒_; Universal; IUniversal; Stable) private variable @@ -204,9 +204,6 @@ module FurtherCorollaries {_<_ : Rel A r} (P : Pred A ℓ) where CounterExample : Pred A ℓ → Pred A ℓ CounterExample P = ¬_ ∘ P -Stable : Pred A ℓ → Set _ -Stable P = ∀ {x} → Negation.Stable (P x) - module _ {_<_ : Rel A r} {P : Pred A ℓ} (stable : Stable P) where open FurtherCorollaries {_<_ = _<_} (CounterExample P) @@ -214,10 +211,10 @@ module _ {_<_ : Rel A r} {P : Pred A ℓ} (stable : Stable P) where renaming (Acc⇒Descent⁺ to NoSmallestCounterExample) acc⇒noSmallestCounterExample : NoSmallestCounterExample → ∀[ Acc _<_ ⇒ P ] - acc⇒noSmallestCounterExample noSmallest = stable ∘ (acc⇒noInfiniteDescent⁺ noSmallest) + acc⇒noSmallestCounterExample noSmallest = stable _ ∘ acc⇒noInfiniteDescent⁺ noSmallest wf⇒noSmallestCounterExample : WellFounded _<_ → NoSmallestCounterExample → ∀[ P ] - wf⇒noSmallestCounterExample wf noSmallest = stable (wf⇒noInfiniteDescent⁺ noSmallest wf) + wf⇒noSmallestCounterExample wf noSmallest = stable _ (wf⇒noInfiniteDescent⁺ noSmallest wf) ------------------------------------------------------------------------ -- Exports diff --git a/src/Relation/Unary.agda b/src/Relation/Unary.agda index f6d7cb13b0..f3cadfb4d2 100644 --- a/src/Relation/Unary.agda +++ b/src/Relation/Unary.agda @@ -15,7 +15,7 @@ open import Data.Sum.Base using (_⊎_; [_,_]) open import Function.Base using (_∘_; _|>_) open import Level using (Level; _⊔_; 0ℓ; suc; Lift) open import Relation.Nullary.Decidable.Core using (Dec; True) -open import Relation.Nullary.Negation.Core using (¬_) +open import Relation.Nullary.Negation.Core as Negation using (¬_) open import Relation.Binary.PropositionalEquality.Core using (_≡_) private @@ -162,6 +162,11 @@ IUniversal P = ∀ {x} → x ∈ P syntax IUniversal P = ∀[ P ] +-- Stability - instances of P are stable wrt double negation + +Stable : Pred A ℓ → Set _ +Stable P = ∀ x → Negation.Stable (P x) + -- Decidability - it is possible to determine if an arbitrary element -- satisfies P. From 96175ed2b17ac745a613a0d22008a77453cac029 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 15 Oct 2023 11:54:00 +0100 Subject: [PATCH 13/57] refactoring; remove explicit definition of `CounterExample` --- src/Induction/NoInfiniteDescent.agda | 13 +++++-------- 1 file changed, 5 insertions(+), 8 deletions(-) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 636e84c898..54f2b89710 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -14,13 +14,15 @@ open import Data.Nat.Base as ℕ using (ℕ; zero; suc) open import Data.Nat.Properties as ℕ open import Data.Product.Base using (_,_; proj₁; ∃-syntax; _×_) open import Function.Base using (_∘_) -open import Induction.WellFounded using (WellFounded; Acc; acc; acc-inverse; module Some) +open import Induction.WellFounded + using (WellFounded; Acc; acc; acc-inverse; module Some) open import Level using (Level) open import Relation.Binary.Core using (Rel) open import Relation.Binary.Construct.Closure.Transitive open import Relation.Binary.PropositionalEquality.Core open import Relation.Nullary.Negation.Core as Negation using (¬_) -open import Relation.Unary using (Pred; _∩_; _⇒_; Universal; IUniversal; Stable) +open import Relation.Unary + using (Pred; ∁; _∩_; _⇒_; Universal; IUniversal; Stable) private variable @@ -199,14 +201,9 @@ module FurtherCorollaries {_<_ : Rel A r} (P : Pred A ℓ) where ------------------------------------------------------------------------ -- Finally: the (classical) 'no smallest counterexample' principle itself --- these belong in Relation.Unary - -CounterExample : Pred A ℓ → Pred A ℓ -CounterExample P = ¬_ ∘ P - module _ {_<_ : Rel A r} {P : Pred A ℓ} (stable : Stable P) where - open FurtherCorollaries {_<_ = _<_} (CounterExample P) + open FurtherCorollaries {_<_ = _<_} (∁ P) using (acc⇒noInfiniteDescent⁺; wf⇒noInfiniteDescent⁺) renaming (Acc⇒Descent⁺ to NoSmallestCounterExample) From e37fe2a3e563498f1207f2a2418942fc006da055 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 15 Oct 2023 11:55:55 +0100 Subject: [PATCH 14/57] refactoring; rename qualified `import` --- src/Relation/Unary.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Relation/Unary.agda b/src/Relation/Unary.agda index f3cadfb4d2..ce121aa755 100644 --- a/src/Relation/Unary.agda +++ b/src/Relation/Unary.agda @@ -15,7 +15,7 @@ open import Data.Sum.Base using (_⊎_; [_,_]) open import Function.Base using (_∘_; _|>_) open import Level using (Level; _⊔_; 0ℓ; suc; Lift) open import Relation.Nullary.Decidable.Core using (Dec; True) -open import Relation.Nullary.Negation.Core as Negation using (¬_) +open import Relation.Nullary.Negation.Core as Nullary using (¬_) open import Relation.Binary.PropositionalEquality.Core using (_≡_) private @@ -165,7 +165,7 @@ syntax IUniversal P = ∀[ P ] -- Stability - instances of P are stable wrt double negation Stable : Pred A ℓ → Set _ -Stable P = ∀ x → Negation.Stable (P x) +Stable P = ∀ x → Nullary.Stable (P x) -- Decidability - it is possible to determine if an arbitrary element -- satisfies P. From 596aac22fd4f09a584ba1712a38d4a1423de6170 Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Wed, 11 Oct 2023 03:00:44 +0100 Subject: [PATCH 15/57] [fixes #2130] Moving `Properties.HeytingAlgebra` from `Relation.Binary` to `Relation.Binary.Lattice` (#2131) --- CHANGELOG.md | 1 + .../Lattice/Properties/HeytingAlgebra.agda | 199 ++++++++++++++++++ .../Binary/Properties/HeytingAlgebra.agda | 195 +---------------- 3 files changed, 208 insertions(+), 187 deletions(-) create mode 100644 src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index 4f65bf45c9..d32dcce7a9 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1253,6 +1253,7 @@ Deprecated modules 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.HeytingAlgebra.agda ↦ Relation.Binary.Lattice.Properties.HeytingAlgebra.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 diff --git a/src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda b/src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda new file mode 100644 index 0000000000..e7090dc6c5 --- /dev/null +++ b/src/Relation/Binary/Lattice/Properties/HeytingAlgebra.agda @@ -0,0 +1,199 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Properties satisfied by Heyting Algebra +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +open import Relation.Binary.Lattice + +module Relation.Binary.Lattice.Properties.HeytingAlgebra + {c ℓ₁ ℓ₂} (L : HeytingAlgebra c ℓ₁ ℓ₂) where + +open HeytingAlgebra L + +open import Algebra.Core +open import Algebra.Definitions _≈_ +open import Data.Product.Base using (_,_) +open import Function.Base using (_$_; flip; _∘_) +open import Level using (_⊔_) +open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_) +import Relation.Binary.Reasoning.PartialOrder as POR +open import Relation.Binary.Lattice.Properties.MeetSemilattice meetSemilattice +open import Relation.Binary.Lattice.Properties.JoinSemilattice joinSemilattice +import Relation.Binary.Lattice.Properties.BoundedMeetSemilattice boundedMeetSemilattice as BM +open import Relation.Binary.Lattice.Properties.Lattice lattice +open import Relation.Binary.Lattice.Properties.BoundedLattice boundedLattice +import Relation.Binary.Reasoning.Setoid as EqReasoning + +------------------------------------------------------------------------ +-- Useful lemmas + +⇨-eval : ∀ {x y} → (x ⇨ y) ∧ x ≤ y +⇨-eval {x} {y} = transpose-∧ refl + +swap-transpose-⇨ : ∀ {x y w} → x ∧ w ≤ y → w ≤ x ⇨ y +swap-transpose-⇨ x∧w≤y = transpose-⇨ $ trans (reflexive $ ∧-comm _ _) x∧w≤y + +------------------------------------------------------------------------ +-- Properties of exponential + +⇨-unit : ∀ {x} → x ⇨ x ≈ ⊤ +⇨-unit = antisym (maximum _) (transpose-⇨ $ reflexive $ BM.identityˡ _) + +y≤x⇨y : ∀ {x y} → y ≤ x ⇨ y +y≤x⇨y = transpose-⇨ (x∧y≤x _ _) + +⇨-drop : ∀ {x y} → (x ⇨ y) ∧ y ≈ y +⇨-drop = antisym (x∧y≤y _ _) (∧-greatest y≤x⇨y refl) + +⇨-app : ∀ {x y} → (x ⇨ y) ∧ x ≈ y ∧ x +⇨-app = antisym (∧-greatest ⇨-eval (x∧y≤y _ _)) (∧-monotonic y≤x⇨y refl) + +⇨ʳ-covariant : ∀ {x} → (x ⇨_) Preserves _≤_ ⟶ _≤_ +⇨ʳ-covariant y≤z = transpose-⇨ (trans ⇨-eval y≤z) + +⇨ˡ-contravariant : ∀ {x} → (_⇨ x) Preserves (flip _≤_) ⟶ _≤_ +⇨ˡ-contravariant z≤y = transpose-⇨ (trans (∧-monotonic refl z≤y) ⇨-eval) + +⇨-relax : _⇨_ Preserves₂ (flip _≤_) ⟶ _≤_ ⟶ _≤_ +⇨-relax {x} {y} {u} {v} y≤x u≤v = begin + x ⇨ u ≤⟨ ⇨ʳ-covariant u≤v ⟩ + x ⇨ v ≤⟨ ⇨ˡ-contravariant y≤x ⟩ + y ⇨ v ∎ + where open POR poset + +⇨-cong : _⇨_ Preserves₂ _≈_ ⟶ _≈_ ⟶ _≈_ +⇨-cong x≈y u≈v = antisym (⇨-relax (reflexive $ Eq.sym x≈y) (reflexive u≈v)) + (⇨-relax (reflexive x≈y) (reflexive $ Eq.sym u≈v)) + +⇨-applyˡ : ∀ {w x y} → w ≤ x → (x ⇨ y) ∧ w ≤ y +⇨-applyˡ = transpose-∧ ∘ ⇨ˡ-contravariant + +⇨-applyʳ : ∀ {w x y} → w ≤ x → w ∧ (x ⇨ y) ≤ y +⇨-applyʳ w≤x = trans (reflexive (∧-comm _ _)) (⇨-applyˡ w≤x) + +⇨-curry : ∀ {x y z} → x ∧ y ⇨ z ≈ x ⇨ y ⇨ z +⇨-curry = antisym (transpose-⇨ $ transpose-⇨ $ trans (reflexive $ ∧-assoc _ _ _) ⇨-eval) + (transpose-⇨ $ trans (reflexive $ Eq.sym $ ∧-assoc _ _ _) + (transpose-∧ $ ⇨-applyˡ refl)) + +------------------------------------------------------------------------ +-- Various proofs of distributivity + +∧-distribˡ-∨-≤ : ∀ x y z → x ∧ (y ∨ z) ≤ x ∧ y ∨ x ∧ z +∧-distribˡ-∨-≤ x y z = trans (reflexive $ ∧-comm _ _) + (transpose-∧ $ ∨-least (swap-transpose-⇨ (x≤x∨y _ _)) $ swap-transpose-⇨ (y≤x∨y _ _)) + +∧-distribˡ-∨-≥ : ∀ x y z → x ∧ y ∨ x ∧ z ≤ x ∧ (y ∨ z) +∧-distribˡ-∨-≥ x y z = let + x∧y≤x , x∧y≤y , _ = infimum x y + x∧z≤x , x∧z≤z , _ = infimum x z + y≤y∨z , z≤y∨z , _ = supremum y z + in ∧-greatest (∨-least x∧y≤x x∧z≤x) + (∨-least (trans x∧y≤y y≤y∨z) (trans x∧z≤z z≤y∨z)) + +∧-distribˡ-∨ : _∧_ DistributesOverˡ _∨_ +∧-distribˡ-∨ x y z = antisym (∧-distribˡ-∨-≤ x y z) (∧-distribˡ-∨-≥ x y z) + +⇨-distribˡ-∧-≤ : ∀ x y z → x ⇨ y ∧ z ≤ (x ⇨ y) ∧ (x ⇨ z) +⇨-distribˡ-∧-≤ x y z = let + y∧z≤y , y∧z≤z , _ = infimum y z + in ∧-greatest (transpose-⇨ $ trans ⇨-eval y∧z≤y) + (transpose-⇨ $ trans ⇨-eval y∧z≤z) + +⇨-distribˡ-∧-≥ : ∀ x y z → (x ⇨ y) ∧ (x ⇨ z) ≤ x ⇨ y ∧ z +⇨-distribˡ-∧-≥ x y z = transpose-⇨ (begin + (((x ⇨ y) ∧ (x ⇨ z)) ∧ x) ≈⟨ ∧-cong Eq.refl $ Eq.sym $ ∧-idempotent _ ⟩ + (((x ⇨ y) ∧ (x ⇨ z)) ∧ x ∧ x) ≈⟨ Eq.sym $ ∧-assoc _ _ _ ⟩ + (((x ⇨ y) ∧ (x ⇨ z)) ∧ x) ∧ x ≈⟨ ∧-cong (∧-assoc _ _ _) Eq.refl ⟩ + (((x ⇨ y) ∧ (x ⇨ z) ∧ x) ∧ x) ≈⟨ ∧-cong (∧-cong Eq.refl $ ∧-comm _ _) Eq.refl ⟩ + (((x ⇨ y) ∧ x ∧ (x ⇨ z)) ∧ x) ≈⟨ ∧-cong (Eq.sym $ ∧-assoc _ _ _) Eq.refl ⟩ + (((x ⇨ y) ∧ x) ∧ (x ⇨ z)) ∧ x ≈⟨ ∧-assoc _ _ _ ⟩ + (((x ⇨ y) ∧ x) ∧ (x ⇨ z) ∧ x) ≤⟨ ∧-monotonic ⇨-eval ⇨-eval ⟩ + y ∧ z ∎) + where open POR poset + +⇨-distribˡ-∧ : _⇨_ DistributesOverˡ _∧_ +⇨-distribˡ-∧ x y z = antisym (⇨-distribˡ-∧-≤ x y z) (⇨-distribˡ-∧-≥ x y z) + +⇨-distribˡ-∨-∧-≤ : ∀ x y z → x ∨ y ⇨ z ≤ (x ⇨ z) ∧ (y ⇨ z) +⇨-distribˡ-∨-∧-≤ x y z = let x≤x∨y , y≤x∨y , _ = supremum x y + in ∧-greatest (transpose-⇨ $ trans (∧-monotonic refl x≤x∨y) ⇨-eval) + (transpose-⇨ $ trans (∧-monotonic refl y≤x∨y) ⇨-eval) + +⇨-distribˡ-∨-∧-≥ : ∀ x y z → (x ⇨ z) ∧ (y ⇨ z) ≤ x ∨ y ⇨ z +⇨-distribˡ-∨-∧-≥ x y z = transpose-⇨ (trans (reflexive $ ∧-distribˡ-∨ _ _ _) + (∨-least (trans (transpose-∧ (x∧y≤x _ _)) refl) + (trans (transpose-∧ (x∧y≤y _ _)) refl))) + +⇨-distribˡ-∨-∧ : ∀ x y z → x ∨ y ⇨ z ≈ (x ⇨ z) ∧ (y ⇨ z) +⇨-distribˡ-∨-∧ x y z = antisym (⇨-distribˡ-∨-∧-≤ x y z) (⇨-distribˡ-∨-∧-≥ x y z) + +------------------------------------------------------------------------ +-- Heyting algebras are distributive lattices + +isDistributiveLattice : IsDistributiveLattice _≈_ _≤_ _∨_ _∧_ +isDistributiveLattice = record + { isLattice = isLattice + ; ∧-distribˡ-∨ = ∧-distribˡ-∨ + } + +distributiveLattice : DistributiveLattice _ _ _ +distributiveLattice = record + { isDistributiveLattice = isDistributiveLattice + } + +------------------------------------------------------------------------ +-- Heyting algebras can define pseudo-complement + +infix 8 ¬_ + +¬_ : Op₁ Carrier +¬ x = x ⇨ ⊥ + +x≤¬¬x : ∀ x → x ≤ ¬ ¬ x +x≤¬¬x x = transpose-⇨ (trans (reflexive (∧-comm _ _)) ⇨-eval) + +------------------------------------------------------------------------ +-- De-Morgan laws + +de-morgan₁ : ∀ x y → ¬ (x ∨ y) ≈ ¬ x ∧ ¬ y +de-morgan₁ x y = ⇨-distribˡ-∨-∧ _ _ _ + +de-morgan₂-≤ : ∀ x y → ¬ (x ∧ y) ≤ ¬ ¬ (¬ x ∨ ¬ y) +de-morgan₂-≤ x y = transpose-⇨ $ begin + ¬ (x ∧ y) ∧ ¬ (¬ x ∨ ¬ y) ≈⟨ ∧-cong ⇨-curry (de-morgan₁ _ _) ⟩ + (x ⇨ ¬ y) ∧ ¬ ¬ x ∧ ¬ ¬ y ≈⟨ ∧-cong Eq.refl (∧-comm _ _) ⟩ + (x ⇨ ¬ y) ∧ ¬ ¬ y ∧ ¬ ¬ x ≈⟨ Eq.sym $ ∧-assoc _ _ _ ⟩ + ((x ⇨ ¬ y) ∧ ¬ ¬ y) ∧ ¬ ¬ x ≤⟨ ⇨-applyʳ $ transpose-⇨ $ + begin + ((x ⇨ ¬ y) ∧ ¬ ¬ y) ∧ x ≈⟨ ∧-cong (∧-comm _ _) Eq.refl ⟩ + ((¬ ¬ y) ∧ (x ⇨ ¬ y)) ∧ x ≈⟨ ∧-assoc _ _ _ ⟩ + (¬ ¬ y) ∧ (x ⇨ ¬ y) ∧ x ≤⟨ ∧-monotonic refl ⇨-eval ⟩ + ¬ ¬ y ∧ ¬ y ≤⟨ ⇨-eval ⟩ + ⊥ ∎ ⟩ + ⊥ ∎ + where open POR poset + +de-morgan₂-≥ : ∀ x y → ¬ ¬ (¬ x ∨ ¬ y) ≤ ¬ (x ∧ y) +de-morgan₂-≥ x y = transpose-⇨ $ ⇨-applyˡ $ transpose-⇨ $ begin + (x ∧ y) ∧ (¬ x ∨ ¬ y) ≈⟨ ∧-distribˡ-∨ _ _ _ ⟩ + (x ∧ y) ∧ ¬ x ∨ (x ∧ y) ∧ ¬ y ≤⟨ ∨-monotonic (⇨-applyʳ (x∧y≤x _ _)) + (⇨-applyʳ (x∧y≤y _ _)) ⟩ + ⊥ ∨ ⊥ ≈⟨ ∨-idempotent _ ⟩ + ⊥ ∎ + where open POR poset + +de-morgan₂ : ∀ x y → ¬ (x ∧ y) ≈ ¬ ¬ (¬ x ∨ ¬ y) +de-morgan₂ x y = antisym (de-morgan₂-≤ x y) (de-morgan₂-≥ x y) + +weak-lem : ∀ {x} → ¬ ¬ (¬ x ∨ x) ≈ ⊤ +weak-lem {x} = begin + ¬ ¬ (¬ x ∨ x) ≈⟨ ⇨-cong (de-morgan₁ _ _) Eq.refl ⟩ + ¬ (¬ ¬ x ∧ ¬ x) ≈⟨ ⇨-cong ⇨-app Eq.refl ⟩ + ⊥ ∧ (x ⇨ ⊥) ⇨ ⊥ ≈⟨ ⇨-cong (∧-zeroˡ _) Eq.refl ⟩ + ⊥ ⇨ ⊥ ≈⟨ ⇨-unit ⟩ + ⊤ ∎ + where open EqReasoning setoid diff --git a/src/Relation/Binary/Properties/HeytingAlgebra.agda b/src/Relation/Binary/Properties/HeytingAlgebra.agda index 87cb01899d..7ec65c0675 100644 --- a/src/Relation/Binary/Properties/HeytingAlgebra.agda +++ b/src/Relation/Binary/Properties/HeytingAlgebra.agda @@ -1,199 +1,20 @@ ------------------------------------------------------------------------ -- The Agda standard library -- --- Properties satisfied by Heyting Algebra +-- This module is DEPRECATED. Please use +-- `Relation.Binary.Lattice.Properties.HeytingAlgebra` instead. ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} -open import Relation.Binary.Lattice +open import Relation.Binary.Lattice using (HeytingAlgebra) module Relation.Binary.Properties.HeytingAlgebra {c ℓ₁ ℓ₂} (L : HeytingAlgebra c ℓ₁ ℓ₂) where -open HeytingAlgebra L +open import Relation.Binary.Lattice.Properties.HeytingAlgebra L public -open import Algebra.Core -open import Algebra.Definitions _≈_ -open import Data.Product.Base using (_,_) -open import Function.Base using (_$_; flip; _∘_) -open import Level using (_⊔_) -open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_) -import Relation.Binary.Reasoning.PartialOrder as POR -open import Relation.Binary.Lattice.Properties.MeetSemilattice meetSemilattice -open import Relation.Binary.Lattice.Properties.JoinSemilattice joinSemilattice -import Relation.Binary.Lattice.Properties.BoundedMeetSemilattice boundedMeetSemilattice as BM -open import Relation.Binary.Lattice.Properties.Lattice lattice -open import Relation.Binary.Lattice.Properties.BoundedLattice boundedLattice -import Relation.Binary.Reasoning.Setoid as EqReasoning - ------------------------------------------------------------------------- --- Useful lemmas - -⇨-eval : ∀ {x y} → (x ⇨ y) ∧ x ≤ y -⇨-eval {x} {y} = transpose-∧ refl - -swap-transpose-⇨ : ∀ {x y w} → x ∧ w ≤ y → w ≤ x ⇨ y -swap-transpose-⇨ x∧w≤y = transpose-⇨ $ trans (reflexive $ ∧-comm _ _) x∧w≤y - ------------------------------------------------------------------------- --- Properties of exponential - -⇨-unit : ∀ {x} → x ⇨ x ≈ ⊤ -⇨-unit = antisym (maximum _) (transpose-⇨ $ reflexive $ BM.identityˡ _) - -y≤x⇨y : ∀ {x y} → y ≤ x ⇨ y -y≤x⇨y = transpose-⇨ (x∧y≤x _ _) - -⇨-drop : ∀ {x y} → (x ⇨ y) ∧ y ≈ y -⇨-drop = antisym (x∧y≤y _ _) (∧-greatest y≤x⇨y refl) - -⇨-app : ∀ {x y} → (x ⇨ y) ∧ x ≈ y ∧ x -⇨-app = antisym (∧-greatest ⇨-eval (x∧y≤y _ _)) (∧-monotonic y≤x⇨y refl) - -⇨ʳ-covariant : ∀ {x} → (x ⇨_) Preserves _≤_ ⟶ _≤_ -⇨ʳ-covariant y≤z = transpose-⇨ (trans ⇨-eval y≤z) - -⇨ˡ-contravariant : ∀ {x} → (_⇨ x) Preserves (flip _≤_) ⟶ _≤_ -⇨ˡ-contravariant z≤y = transpose-⇨ (trans (∧-monotonic refl z≤y) ⇨-eval) - -⇨-relax : _⇨_ Preserves₂ (flip _≤_) ⟶ _≤_ ⟶ _≤_ -⇨-relax {x} {y} {u} {v} y≤x u≤v = begin - x ⇨ u ≤⟨ ⇨ʳ-covariant u≤v ⟩ - x ⇨ v ≤⟨ ⇨ˡ-contravariant y≤x ⟩ - y ⇨ v ∎ - where open POR poset - -⇨-cong : _⇨_ Preserves₂ _≈_ ⟶ _≈_ ⟶ _≈_ -⇨-cong x≈y u≈v = antisym (⇨-relax (reflexive $ Eq.sym x≈y) (reflexive u≈v)) - (⇨-relax (reflexive x≈y) (reflexive $ Eq.sym u≈v)) - -⇨-applyˡ : ∀ {w x y} → w ≤ x → (x ⇨ y) ∧ w ≤ y -⇨-applyˡ = transpose-∧ ∘ ⇨ˡ-contravariant - -⇨-applyʳ : ∀ {w x y} → w ≤ x → w ∧ (x ⇨ y) ≤ y -⇨-applyʳ w≤x = trans (reflexive (∧-comm _ _)) (⇨-applyˡ w≤x) - -⇨-curry : ∀ {x y z} → x ∧ y ⇨ z ≈ x ⇨ y ⇨ z -⇨-curry = antisym (transpose-⇨ $ transpose-⇨ $ trans (reflexive $ ∧-assoc _ _ _) ⇨-eval) - (transpose-⇨ $ trans (reflexive $ Eq.sym $ ∧-assoc _ _ _) - (transpose-∧ $ ⇨-applyˡ refl)) - ------------------------------------------------------------------------- --- Various proofs of distributivity - -∧-distribˡ-∨-≤ : ∀ x y z → x ∧ (y ∨ z) ≤ x ∧ y ∨ x ∧ z -∧-distribˡ-∨-≤ x y z = trans (reflexive $ ∧-comm _ _) - (transpose-∧ $ ∨-least (swap-transpose-⇨ (x≤x∨y _ _)) $ swap-transpose-⇨ (y≤x∨y _ _)) - -∧-distribˡ-∨-≥ : ∀ x y z → x ∧ y ∨ x ∧ z ≤ x ∧ (y ∨ z) -∧-distribˡ-∨-≥ x y z = let - x∧y≤x , x∧y≤y , _ = infimum x y - x∧z≤x , x∧z≤z , _ = infimum x z - y≤y∨z , z≤y∨z , _ = supremum y z - in ∧-greatest (∨-least x∧y≤x x∧z≤x) - (∨-least (trans x∧y≤y y≤y∨z) (trans x∧z≤z z≤y∨z)) - -∧-distribˡ-∨ : _∧_ DistributesOverˡ _∨_ -∧-distribˡ-∨ x y z = antisym (∧-distribˡ-∨-≤ x y z) (∧-distribˡ-∨-≥ x y z) - -⇨-distribˡ-∧-≤ : ∀ x y z → x ⇨ y ∧ z ≤ (x ⇨ y) ∧ (x ⇨ z) -⇨-distribˡ-∧-≤ x y z = let - y∧z≤y , y∧z≤z , _ = infimum y z - in ∧-greatest (transpose-⇨ $ trans ⇨-eval y∧z≤y) - (transpose-⇨ $ trans ⇨-eval y∧z≤z) - -⇨-distribˡ-∧-≥ : ∀ x y z → (x ⇨ y) ∧ (x ⇨ z) ≤ x ⇨ y ∧ z -⇨-distribˡ-∧-≥ x y z = transpose-⇨ (begin - (((x ⇨ y) ∧ (x ⇨ z)) ∧ x) ≈⟨ ∧-cong Eq.refl $ Eq.sym $ ∧-idempotent _ ⟩ - (((x ⇨ y) ∧ (x ⇨ z)) ∧ x ∧ x) ≈⟨ Eq.sym $ ∧-assoc _ _ _ ⟩ - (((x ⇨ y) ∧ (x ⇨ z)) ∧ x) ∧ x ≈⟨ ∧-cong (∧-assoc _ _ _) Eq.refl ⟩ - (((x ⇨ y) ∧ (x ⇨ z) ∧ x) ∧ x) ≈⟨ ∧-cong (∧-cong Eq.refl $ ∧-comm _ _) Eq.refl ⟩ - (((x ⇨ y) ∧ x ∧ (x ⇨ z)) ∧ x) ≈⟨ ∧-cong (Eq.sym $ ∧-assoc _ _ _) Eq.refl ⟩ - (((x ⇨ y) ∧ x) ∧ (x ⇨ z)) ∧ x ≈⟨ ∧-assoc _ _ _ ⟩ - (((x ⇨ y) ∧ x) ∧ (x ⇨ z) ∧ x) ≤⟨ ∧-monotonic ⇨-eval ⇨-eval ⟩ - y ∧ z ∎) - where open POR poset - -⇨-distribˡ-∧ : _⇨_ DistributesOverˡ _∧_ -⇨-distribˡ-∧ x y z = antisym (⇨-distribˡ-∧-≤ x y z) (⇨-distribˡ-∧-≥ x y z) - -⇨-distribˡ-∨-∧-≤ : ∀ x y z → x ∨ y ⇨ z ≤ (x ⇨ z) ∧ (y ⇨ z) -⇨-distribˡ-∨-∧-≤ x y z = let x≤x∨y , y≤x∨y , _ = supremum x y - in ∧-greatest (transpose-⇨ $ trans (∧-monotonic refl x≤x∨y) ⇨-eval) - (transpose-⇨ $ trans (∧-monotonic refl y≤x∨y) ⇨-eval) - -⇨-distribˡ-∨-∧-≥ : ∀ x y z → (x ⇨ z) ∧ (y ⇨ z) ≤ x ∨ y ⇨ z -⇨-distribˡ-∨-∧-≥ x y z = transpose-⇨ (trans (reflexive $ ∧-distribˡ-∨ _ _ _) - (∨-least (trans (transpose-∧ (x∧y≤x _ _)) refl) - (trans (transpose-∧ (x∧y≤y _ _)) refl))) - -⇨-distribˡ-∨-∧ : ∀ x y z → x ∨ y ⇨ z ≈ (x ⇨ z) ∧ (y ⇨ z) -⇨-distribˡ-∨-∧ x y z = antisym (⇨-distribˡ-∨-∧-≤ x y z) (⇨-distribˡ-∨-∧-≥ x y z) - ------------------------------------------------------------------------- --- Heyting algebras are distributive lattices - -isDistributiveLattice : IsDistributiveLattice _≈_ _≤_ _∨_ _∧_ -isDistributiveLattice = record - { isLattice = isLattice - ; ∧-distribˡ-∨ = ∧-distribˡ-∨ - } - -distributiveLattice : DistributiveLattice _ _ _ -distributiveLattice = record - { isDistributiveLattice = isDistributiveLattice - } - ------------------------------------------------------------------------- --- Heyting algebras can define pseudo-complement - -infix 8 ¬_ - -¬_ : Op₁ Carrier -¬ x = x ⇨ ⊥ - -x≤¬¬x : ∀ x → x ≤ ¬ ¬ x -x≤¬¬x x = transpose-⇨ (trans (reflexive (∧-comm _ _)) ⇨-eval) - ------------------------------------------------------------------------- --- De-Morgan laws - -de-morgan₁ : ∀ x y → ¬ (x ∨ y) ≈ ¬ x ∧ ¬ y -de-morgan₁ x y = ⇨-distribˡ-∨-∧ _ _ _ - -de-morgan₂-≤ : ∀ x y → ¬ (x ∧ y) ≤ ¬ ¬ (¬ x ∨ ¬ y) -de-morgan₂-≤ x y = transpose-⇨ $ begin - ¬ (x ∧ y) ∧ ¬ (¬ x ∨ ¬ y) ≈⟨ ∧-cong ⇨-curry (de-morgan₁ _ _) ⟩ - (x ⇨ ¬ y) ∧ ¬ ¬ x ∧ ¬ ¬ y ≈⟨ ∧-cong Eq.refl (∧-comm _ _) ⟩ - (x ⇨ ¬ y) ∧ ¬ ¬ y ∧ ¬ ¬ x ≈⟨ Eq.sym $ ∧-assoc _ _ _ ⟩ - ((x ⇨ ¬ y) ∧ ¬ ¬ y) ∧ ¬ ¬ x ≤⟨ ⇨-applyʳ $ transpose-⇨ $ - begin - ((x ⇨ ¬ y) ∧ ¬ ¬ y) ∧ x ≈⟨ ∧-cong (∧-comm _ _) Eq.refl ⟩ - ((¬ ¬ y) ∧ (x ⇨ ¬ y)) ∧ x ≈⟨ ∧-assoc _ _ _ ⟩ - (¬ ¬ y) ∧ (x ⇨ ¬ y) ∧ x ≤⟨ ∧-monotonic refl ⇨-eval ⟩ - ¬ ¬ y ∧ ¬ y ≤⟨ ⇨-eval ⟩ - ⊥ ∎ ⟩ - ⊥ ∎ - where open POR poset - -de-morgan₂-≥ : ∀ x y → ¬ ¬ (¬ x ∨ ¬ y) ≤ ¬ (x ∧ y) -de-morgan₂-≥ x y = transpose-⇨ $ ⇨-applyˡ $ transpose-⇨ $ begin - (x ∧ y) ∧ (¬ x ∨ ¬ y) ≈⟨ ∧-distribˡ-∨ _ _ _ ⟩ - (x ∧ y) ∧ ¬ x ∨ (x ∧ y) ∧ ¬ y ≤⟨ ∨-monotonic (⇨-applyʳ (x∧y≤x _ _)) - (⇨-applyʳ (x∧y≤y _ _)) ⟩ - ⊥ ∨ ⊥ ≈⟨ ∨-idempotent _ ⟩ - ⊥ ∎ - where open POR poset - -de-morgan₂ : ∀ x y → ¬ (x ∧ y) ≈ ¬ ¬ (¬ x ∨ ¬ y) -de-morgan₂ x y = antisym (de-morgan₂-≤ x y) (de-morgan₂-≥ x y) - -weak-lem : ∀ {x} → ¬ ¬ (¬ x ∨ x) ≈ ⊤ -weak-lem {x} = begin - ¬ ¬ (¬ x ∨ x) ≈⟨ ⇨-cong (de-morgan₁ _ _) Eq.refl ⟩ - ¬ (¬ ¬ x ∧ ¬ x) ≈⟨ ⇨-cong ⇨-app Eq.refl ⟩ - ⊥ ∧ (x ⇨ ⊥) ⇨ ⊥ ≈⟨ ⇨-cong (∧-zeroˡ _) Eq.refl ⟩ - ⊥ ⇨ ⊥ ≈⟨ ⇨-unit ⟩ - ⊤ ∎ - where open EqReasoning setoid +{-# WARNING_ON_IMPORT +"Relation.Binary.Properties.HeytingAlgebra was deprecated in v2.0. +Use Relation.Binary.Lattice.Properties.HeytingAlgebra instead." +#-} From e06c02aae2f639659feb5cc159c604c262c70cfc Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Wed, 11 Oct 2023 03:01:16 +0100 Subject: [PATCH 16/57] [fixes #2127] Fixes #1930 `import` bug (#2128) --- README.agda | 5 +++++ README/Inspect.agda | 9 +++------ 2 files changed, 8 insertions(+), 6 deletions(-) diff --git a/README.agda b/README.agda index 97b8ab3903..ced0f11596 100644 --- a/README.agda +++ b/README.agda @@ -237,6 +237,11 @@ import README.Debug.Trace import README.Nary +-- Explaining the inspect idiom: use case, equivalent handwritten +-- auxiliary definitions, and implementation details. + +import README.Inspect + -- Explaining how to use the automatic solvers import README.Tactic.MonoidSolver diff --git a/README/Inspect.agda b/README/Inspect.agda index 9f40a8d65a..2199a499c8 100644 --- a/README/Inspect.agda +++ b/README/Inspect.agda @@ -1,22 +1,19 @@ ------------------------------------------------------------------------ -- The Agda standard library -- --- This module is DEPRECATED. +-- Explaining how to use the inspect idiom and elaborating on the way +-- it is implemented in the standard library. ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} module README.Inspect where -{-# WARNING_ON_IMPORT -"README.Inspect was deprecated in v2.0." -#-} - open import Data.Nat.Base open import Data.Nat.Properties open import Data.Product.Base using (_×_; _,_) open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) -open import Relation.Binary.PropositionalEquality using (inspect) +open import Relation.Binary.PropositionalEquality using (inspect; [_]) ------------------------------------------------------------------------ -- Using inspect From eee9f58efc94cf2b1bdc78e07d26077293cf4004 Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Thu, 12 Oct 2023 02:52:31 +0100 Subject: [PATCH 17/57] [fixes #1214] Add negated ordering relation symbols systematically to `Relation.Binary.*` (#2095) --- CHANGELOG.md | 43 ++++++++++++----- src/Data/List/Sort/MergeSort.agda | 4 +- src/Relation/Binary/Bundles.agda | 48 ++++++++++--------- .../Binary/Properties/DecTotalOrder.agda | 11 ++--- src/Relation/Binary/Properties/Poset.agda | 12 ++--- .../Binary/Properties/StrictPartialOrder.agda | 2 +- .../Binary/Properties/StrictTotalOrder.agda | 2 +- .../Binary/Properties/TotalOrder.agda | 5 +- 8 files changed, 72 insertions(+), 55 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index d32dcce7a9..545055b93c 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -842,7 +842,6 @@ Non-backwards compatible changes IO.Effectful IO.Instances ``` - ### (Issue #2096) Introduction of flipped relation symbol for `Relation.Binary.Bundles.Preorder` * Previously, the relation symbol `_∼_` was (notationally) symmetric, so that its @@ -856,9 +855,25 @@ Non-backwards compatible changes Therefore, only _declarations_ of `PartialOrder` records will need their field names updating. -* NB (issues #1214 #2098) the corresponding situation regarding the `flip`ped - relation symbols `_≥_`, `_>_` (and their negated versions!) has not (yet) - been addressed. +### (Issue #1214) Reorganisation of the introduction of negated relation symbols under `Relation.Binary` + +* Previously, negated relation symbols `_≰_` (for `Poset`) and `_≮_` (`TotalOrder`) + were introduced in the corresponding `Relation.Binary.Properties` modules, for re-export. + +* Now they are introduced as definitions in the corresponding `Relation.Binary.Bundles`, + together with, for uniformity's sake, an additional negated symbol `_⋦_` for `Preorder`, + with their (obvious) intended semantics: + ```agda + infix 4 _⋦_ _≰_ _≮_ + Preorder._⋦_ : Rel Carrier _ + StrictPartialOrder._≮_ : Rel Carrier _ + ``` + Additionally, `Poset._≰_` is defined by renaming public export of `Preorder._⋦_` + +* Backwards compatibility has been maintained, with deprecated definitions in the + corresponding `Relation.Binary.Properties` modules, and the corresponding client + client module `import`s being adjusted accordingly. + ### Standardisation of `insertAt`/`updateAt`/`removeAt` @@ -1743,11 +1758,6 @@ Deprecated names fromForeign ↦ Foreign.Haskell.Coerce.coerce ``` -* In `Relation.Binary.Bundles.Preorder`: - ``` - _∼_ ↦ _≲_ - ``` - * In `Relation.Binary.Indexed.Heterogeneous.Bundles.Preorder`: ``` _∼_ ↦ _≲_ @@ -1759,6 +1769,15 @@ Deprecated names invPreorder ↦ converse-preorder ``` +* Moved negated relation symbol from `Relation.Binary.Properties.Poset`: + ``` + _≰_ ↦ Relation.Binary.Bundles.Poset._≰_ + ``` + +* Moved negated relation symbol from `Relation.Binary.Properties.TotalOrder`: + ``` + _≮_ ↦ Relation.Binary.Bundles.StrictPartialOrder._≮_ + * In `Relation.Nullary.Decidable.Core`: ``` excluded-middle ↦ ¬¬-excluded-middle @@ -3312,9 +3331,9 @@ Additions to existing modules * Added new proofs to `Relation.Binary.Consequences`: ``` - sym⇒¬-sym : Symmetric _∼_ → Symmetric _≁_ - cotrans⇒¬-trans : Cotransitive _∼_ → Transitive _≁_ - irrefl⇒¬-refl : Reflexive _≈_ → Irreflexive _≈_ _∼_ → Reflexive _≁_ + sym⇒¬-sym : Symmetric _∼_ → Symmetric (¬_ ∘₂ _∼_) + cotrans⇒¬-trans : Cotransitive _∼_ → Transitive (¬_ ∘₂ _∼_) + irrefl⇒¬-refl : Reflexive _≈_ → Irreflexive _≈_ _∼_ → Reflexive (¬_ ∘₂ _∼_) mono₂⇒cong₂ : Symmetric ≈₁ → ≈₁ ⇒ ≤₁ → Antisymmetric ≈₂ ≤₂ → ∀ {f} → f Preserves₂ ≤₁ ⟶ ≤₁ ⟶ ≤₂ → f Preserves₂ ≈₁ ⟶ ≈₁ ⟶ ≈₂ diff --git a/src/Data/List/Sort/MergeSort.agda b/src/Data/List/Sort/MergeSort.agda index 9abab0823e..bded259428 100644 --- a/src/Data/List/Sort/MergeSort.agda +++ b/src/Data/List/Sort/MergeSort.agda @@ -21,7 +21,7 @@ open import Data.List.Relation.Binary.Permutation.Propositional import Data.List.Relation.Binary.Permutation.Propositional.Properties as Perm open import Data.Maybe.Base using (just) open import Relation.Nullary.Decidable using (does) -open import Data.Nat using (_<_; _>_; z_; z ; ≰⇒≥ ) -≮⇒≥ : ∀ {x y} → ¬ (x < y) → y ≤ x +≮⇒≥ : ∀ {x y} → x ≮ y → y ≤ x ≮⇒≥ = ToStrict.≮⇒≥ Eq.sym _≟_ reflexive total diff --git a/src/Relation/Binary/Properties/Poset.agda b/src/Relation/Binary/Properties/Poset.agda index add0e20b28..beb4ddc8aa 100644 --- a/src/Relation/Binary/Properties/Poset.agda +++ b/src/Relation/Binary/Properties/Poset.agda @@ -6,6 +6,7 @@ {-# OPTIONS --cubical-compatible --safe #-} +open import Data.Product.Base using (_,_) open import Function.Base using (flip; _∘_) open import Relation.Binary.Core using (Rel; _Preserves_⟶_) open import Relation.Binary.Bundles using (Poset; StrictPartialOrder) @@ -62,11 +63,6 @@ open Poset ≥-poset public ------------------------------------------------------------------------ -- Negated order -infix 4 _≰_ - -_≰_ : Rel A p₃ -x ≰ y = ¬ (x ≤ y) - ≰-respˡ-≈ : _≰_ Respectsˡ _≈_ ≰-respˡ-≈ x≈y = _∘ ≤-respˡ-≈ (Eq.sym x≈y) @@ -90,7 +86,7 @@ _<_ = ToStrict._<_ } open StrictPartialOrder <-strictPartialOrder public - using ( <-resp-≈; <-respʳ-≈; <-respˡ-≈) + using (_≮_; <-resp-≈; <-respʳ-≈; <-respˡ-≈) renaming ( irrefl to <-irrefl ; asym to <-asym @@ -103,10 +99,10 @@ open StrictPartialOrder <-strictPartialOrder public ≤∧≉⇒< : ∀ {x y} → x ≤ y → x ≉ y → x < y ≤∧≉⇒< = ToStrict.≤∧≉⇒< -<⇒≱ : ∀ {x y} → x < y → ¬ (y ≤ x) +<⇒≱ : ∀ {x y} → x < y → y ≰ x <⇒≱ = ToStrict.<⇒≱ antisym -≤⇒≯ : ∀ {x y} → x ≤ y → ¬ (y < x) +≤⇒≯ : ∀ {x y} → x ≤ y → y ≮ x ≤⇒≯ = ToStrict.≤⇒≯ antisym ------------------------------------------------------------------------ diff --git a/src/Relation/Binary/Properties/StrictPartialOrder.agda b/src/Relation/Binary/Properties/StrictPartialOrder.agda index 7420a2e24f..680323c770 100644 --- a/src/Relation/Binary/Properties/StrictPartialOrder.agda +++ b/src/Relation/Binary/Properties/StrictPartialOrder.agda @@ -17,7 +17,7 @@ import Relation.Binary.Construct.StrictToNonStrict open StrictPartialOrder SPO ------------------------------------------------------------------------ --- The inverse relation is also a strict partial order. +-- The converse relation is also a strict partial order. >-strictPartialOrder : StrictPartialOrder s₁ s₂ s₃ >-strictPartialOrder = EqAndOrd.strictPartialOrder SPO diff --git a/src/Relation/Binary/Properties/StrictTotalOrder.agda b/src/Relation/Binary/Properties/StrictTotalOrder.agda index f8b0062474..e6d34ab8d1 100644 --- a/src/Relation/Binary/Properties/StrictTotalOrder.agda +++ b/src/Relation/Binary/Properties/StrictTotalOrder.agda @@ -18,7 +18,7 @@ import Relation.Binary.Properties.StrictPartialOrder as SPO open import Relation.Binary.Consequences ------------------------------------------------------------------------ --- _<_ - the strict version is a strict total order +-- _<_ - the strict version is a decidable total order decTotalOrder : DecTotalOrder _ _ _ decTotalOrder = record diff --git a/src/Relation/Binary/Properties/TotalOrder.agda b/src/Relation/Binary/Properties/TotalOrder.agda index 540de19905..48f788ab7f 100644 --- a/src/Relation/Binary/Properties/TotalOrder.agda +++ b/src/Relation/Binary/Properties/TotalOrder.agda @@ -59,7 +59,7 @@ open PosetProperties public } open TotalOrder ≥-totalOrder public - using () renaming (total to ≥-total) + using () renaming (total to ≥-total; _≰_ to _≱_) ------------------------------------------------------------------------ -- _<_ - the strict version is a strict partial order @@ -90,8 +90,7 @@ open PosetProperties public open PosetProperties public using - ( _≰_ - ; ≰-respʳ-≈ + ( ≰-respʳ-≈ ; ≰-respˡ-≈ ) From 81a7b289a442c40ceee68b2080f3f4eba6442bc1 Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Thu, 12 Oct 2023 03:05:52 +0100 Subject: [PATCH 18/57] Refactoring (inversion) properties of `_<_` on `Nat`, plus consequences (#2000) --- CHANGELOG.md | 17 +- src/Data/Digit/Properties.agda | 2 +- src/Data/Fin/Base.agda | 44 +++-- src/Data/Fin/Properties.agda | 170 +++++++++--------- src/Data/Integer/Properties.agda | 26 +-- src/Data/List/Properties.agda | 2 +- .../Relation/Unary/Linked/Properties.agda | 2 +- src/Data/Nat.agda | 3 + src/Data/Nat/Base.agda | 36 +++- src/Data/Nat/Binary/Properties.agda | 10 +- src/Data/Nat/DivMod.agda | 14 +- src/Data/Nat/DivMod/WithK.agda | 18 +- src/Data/Nat/Properties.agda | 52 +++--- src/Data/Nat/Properties/Core.agda | 6 +- src/Data/Nat/Show/Properties.agda | 16 +- src/Data/Vec/Base.agda | 2 +- src/Data/Vec/Bounded/Base.agda | 4 +- src/Data/Vec/Properties.agda | 14 +- .../Vec/Relation/Unary/Linked/Properties.agda | 11 +- 19 files changed, 244 insertions(+), 205 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 545055b93c..711fe9b14b 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1276,7 +1276,7 @@ Deprecated modules ### Deprecation of `Data.Nat.Properties.Core` -* The module `Data.Nat.Properties.Core` has been deprecated, and its one entry moved to `Data.Nat.Properties` +* The module `Data.Nat.Properties.Core` has been deprecated, and its one lemma moved to `Data.Nat.Base`, renamed as `s≤s⁻¹` ### Deprecation of `Data.Fin.Substitution.Example` @@ -1543,6 +1543,7 @@ Deprecated names ≤-stepsˡ ↦ m≤n⇒m≤o+n ≤-stepsʳ ↦ m≤n⇒m≤n+o <-step ↦ m (λ()) (λ()) z i≮j i≢j j (i≮j ∘ ℕₚ.≤-pred) (i≢j ∘ suc-injective) (s i≮j i≢j j (i≮j ∘ s m≮n m≢n n>m = tri< (-<- n>m) (m≢n ∘ -[1+-injective) (m≮n ∘ drop‿-<-) <-cmp +[1+ m ] +[1+ n ] with ℕ.<-cmp m n -... | tri< m m≮n m≢n n>m = tri> (m≮n ∘ ℕ.≤-pred ∘ drop‿+<+) (m≢n ∘ +[1+-injective) (+<+ (sm)) +... | tri< m m≮n m≢n n>m = tri> (m≮n ∘ sm)) infix 4 _ j neg-cancel-< { +[1+ m ]} { +[1+ n ]} (-<- n′_ @@ -308,8 +317,9 @@ m ≥′ n = n ≤′ m _>′_ : Rel ℕ 0ℓ m >′ n = n <′ m ------------------------------------------------------------------------- --- Another alternative definition of _≤_ +-- _≤″_: this definition of _≤_ is used for proof-irrelevant ‵DivMod` +-- and is a specialised instance of a general algebraic construction + infix 4 _≤″_ _<″_ _≥″_ _>″_ _≤″_ : (m n : ℕ) → Set @@ -326,10 +336,20 @@ m ≥″ n = n ≤″ m _>″_ : Rel ℕ 0ℓ m >″ n = n <″ m ------------------------------------------------------------------------- --- Another alternative definition of _≤_ +-- Smart constructors of _≤″_ and _<″_ + +pattern ≤″-offset k = less-than-or-equal {k = k} refl +pattern <″-offset k = ≤″-offset k + +-- Smart destructors of _<″_ + +s≤″s⁻¹ : ∀ {m n} → suc m ≤″ suc n → m ≤″ n +s≤″s⁻¹ (≤″-offset k) = ≤″-offset k + +s<″s⁻¹ : ∀ {m n} → suc m <″ suc n → m <″ n +s<″s⁻¹ (<″-offset k) = <″-offset k --- Useful for induction when you have an upper bound. +-- _≤‴_: this definition is useful for induction with an upper bound. data _≤‴_ : ℕ → ℕ → Set where ≤‴-refl : ∀{m} → m ≤‴ m diff --git a/src/Data/Nat/Binary/Properties.agda b/src/Data/Nat/Binary/Properties.agda index 638482131f..6109106b26 100644 --- a/src/Data/Nat/Binary/Properties.agda +++ b/src/Data/Nat/Binary/Properties.agda @@ -15,7 +15,7 @@ open import Algebra.Consequences.Propositional open import Data.Bool.Base using (if_then_else_; Bool; true; false) open import Data.Maybe.Base using (Maybe; just; nothing) open import Data.Nat.Binary.Base -open import Data.Nat as ℕ using (ℕ; z≤n; s≤s) +open import Data.Nat as ℕ using (ℕ; z≤n; s≤s; s 0 → m < m + n m0 = n>0 @@ -727,8 +723,8 @@ m 0 → m < n + m m0 rewrite +-comm n m = m0 m+n≮n : ∀ m n → m + n ≮ n -m+n≮n zero n = n≮n n -m+n≮n (suc m) (suc n) (s Date: Thu, 12 Oct 2023 13:01:14 +0900 Subject: [PATCH 19/57] Bump CI tests to Agda-2.6.4 (#2134) --- .github/workflows/ci-ubuntu.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/ci-ubuntu.yml b/.github/workflows/ci-ubuntu.yml index 85e9fe1365..ebafe70fa8 100644 --- a/.github/workflows/ci-ubuntu.yml +++ b/.github/workflows/ci-ubuntu.yml @@ -74,7 +74,7 @@ jobs: if [[ '${{ github.ref }}' == 'refs/heads/master' \ || '${{ github.base_ref }}' == 'master' ]]; then # Pick Agda version for master - echo "AGDA_COMMIT=tags/v2.6.3" >> $GITHUB_ENV; + echo "AGDA_COMMIT=tags/v2.6.4" >> $GITHUB_ENV; echo "AGDA_HTML_DIR=html" >> $GITHUB_ENV elif [[ '${{ github.ref }}' == 'refs/heads/experimental' \ || '${{ github.base_ref }}' == 'experimental' ]]; then From cb92af1625cfc8514261b497fd27f10c1ba55211 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Thu, 12 Oct 2023 22:42:47 +0900 Subject: [PATCH 20/57] Remove `Algebra.Ordered` (#2133) --- CHANGELOG.md | 7 - src/Algebra/Ordered.agda | 14 - src/Algebra/Ordered/Bundles.agda | 438 ------------------ src/Algebra/Ordered/Structures.agda | 391 ---------------- .../Properties/BoundedJoinSemilattice.agda | 25 +- .../Lattice/Properties/JoinSemilattice.agda | 22 - .../Lattice/Properties/MeetSemilattice.agda | 2 +- 7 files changed, 2 insertions(+), 897 deletions(-) delete mode 100644 src/Algebra/Ordered.agda delete mode 100644 src/Algebra/Ordered/Bundles.agda delete mode 100644 src/Algebra/Ordered/Structures.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index 711fe9b14b..7755753af2 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1832,13 +1832,6 @@ New modules Algebra.Morphism.Construct.Identity ``` -* Ordered algebraic structures (pomonoids, posemigroups, etc.) - ``` - Algebra.Ordered - Algebra.Ordered.Bundles - Algebra.Ordered.Structures - ``` - * 'Optimised' tail-recursive exponentiation properties: ``` Algebra.Properties.Semiring.Exp.TailRecursiveOptimised diff --git a/src/Algebra/Ordered.agda b/src/Algebra/Ordered.agda deleted file mode 100644 index 5eb99aec36..0000000000 --- a/src/Algebra/Ordered.agda +++ /dev/null @@ -1,14 +0,0 @@ ------------------------------------------------------------------------- --- The Agda standard library --- --- Definitions of ordered algebraic structures like promonoids and --- posemigroups (packed in records together with sets, orders, --- operations, etc.) ------------------------------------------------------------------------- - -{-# OPTIONS --cubical-compatible --safe #-} - -module Algebra.Ordered where - -open import Algebra.Ordered.Structures public -open import Algebra.Ordered.Bundles public diff --git a/src/Algebra/Ordered/Bundles.agda b/src/Algebra/Ordered/Bundles.agda deleted file mode 100644 index 896f9d4d1f..0000000000 --- a/src/Algebra/Ordered/Bundles.agda +++ /dev/null @@ -1,438 +0,0 @@ ------------------------------------------------------------------------- --- The Agda standard library --- --- Definitions of ordered algebraic structures like promonoids and --- posemigroups (packed in records together with sets, orders, --- operations, etc.) ------------------------------------------------------------------------- - --- The contents of this module should be accessed via `Algebra.Order`. - -{-# OPTIONS --cubical-compatible --safe #-} - -module Algebra.Ordered.Bundles where - -open import Algebra.Core -open import Algebra.Bundles -open import Algebra.Ordered.Structures -open import Level using (suc; _⊔_) -open import Relation.Binary.Core using (Rel) -open import Relation.Binary.Bundles using (Preorder; Poset) - ------------------------------------------------------------------------- --- Bundles of preordered structures - --- Preordered magmas (promagmas) - -record Promagma c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The preorder. - _∙_ : Op₂ Carrier -- Multiplication. - isPromagma : IsPromagma _≈_ _≤_ _∙_ - - open IsPromagma isPromagma public - - preorder : Preorder c ℓ₁ ℓ₂ - preorder = record { isPreorder = isPreorder } - - magma : Magma c ℓ₁ - magma = record { isMagma = isMagma } - --- Preordered semigroups (prosemigroups) - -record Prosemigroup c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The preorder. - _∙_ : Op₂ Carrier -- Multiplication. - isProsemigroup : IsProsemigroup _≈_ _≤_ _∙_ - - open IsProsemigroup isProsemigroup public - - promagma : Promagma c ℓ₁ ℓ₂ - promagma = record { isPromagma = isPromagma } - - open Promagma promagma public using (preorder; magma) - - semigroup : Semigroup c ℓ₁ - semigroup = record { isSemigroup = isSemigroup } - --- Preordered monoids (promonoids) - -record Promonoid c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The preorder. - _∙_ : Op₂ Carrier -- The monoid multiplication. - ε : Carrier -- The monoid unit. - isPromonoid : IsPromonoid _≈_ _≤_ _∙_ ε - - open IsPromonoid isPromonoid public - - prosemigroup : Prosemigroup c ℓ₁ ℓ₂ - prosemigroup = record { isProsemigroup = isProsemigroup } - - open Prosemigroup prosemigroup public - using (preorder; magma; promagma; semigroup) - - monoid : Monoid c ℓ₁ - monoid = record { isMonoid = isMonoid } - --- Preordered commutative monoids (commutative promonoids) - -record CommutativePromonoid c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The preorder. - _∙_ : Op₂ Carrier -- The monoid multiplication. - ε : Carrier -- The monoid unit. - isCommutativePromonoid : IsCommutativePromonoid _≈_ _≤_ _∙_ ε - - open IsCommutativePromonoid isCommutativePromonoid public - - promonoid : Promonoid c ℓ₁ ℓ₂ - promonoid = record { isPromonoid = isPromonoid } - - open Promonoid promonoid public - using (preorder; magma; promagma; semigroup; prosemigroup; monoid) - - commutativeMonoid : CommutativeMonoid c ℓ₁ - commutativeMonoid = record { isCommutativeMonoid = isCommutativeMonoid } - - open CommutativeMonoid commutativeMonoid public using (commutativeSemigroup) - --- Preordered semirings (prosemirings) - -record Prosemiring c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _*_ - infixl 6 _+_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ - _≤_ : Rel Carrier ℓ₂ - _+_ : Op₂ Carrier - _*_ : Op₂ Carrier - 0# : Carrier - 1# : Carrier - isProsemiring : IsProsemiring _≈_ _≤_ _+_ _*_ 0# 1# - - open IsProsemiring isProsemiring public - - +-commutativePromonoid : CommutativePromonoid c ℓ₁ ℓ₂ - +-commutativePromonoid = record - { isCommutativePromonoid = +-isCommutativePromonoid } - - open CommutativePromonoid +-commutativePromonoid public - using (preorder) - renaming - ( magma to +-magma - ; promagma to +-promagma - ; semigroup to +-semigroup - ; prosemigroup to +-prosemigroup - ; monoid to +-monoid - ; promonoid to +-promonoid - ; commutativeSemigroup to +-commutativeSemigroup - ; commutativeMonoid to +-commutativeMonoid - ) - - *-promonoid : Promonoid c ℓ₁ ℓ₂ - *-promonoid = record { isPromonoid = *-isPromonoid } - - open Promonoid *-promonoid public - using () - renaming - ( magma to *-magma - ; promagma to *-promagma - ; semigroup to *-semigroup - ; prosemigroup to *-prosemigroup - ; monoid to *-monoid - ) - - semiring : Semiring c ℓ₁ - semiring = record { isSemiring = isSemiring } - --- Preordered IdempotentSemiring (IdempotentProsemiring) - -record IdempotentProsemiring c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _*_ - infixl 6 _+_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ - _≤_ : Rel Carrier ℓ₂ - _+_ : Op₂ Carrier - _*_ : Op₂ Carrier - 0# : Carrier - 1# : Carrier - isIdempotentProsemiring : IsIdempotentProsemiring _≈_ _≤_ _+_ _*_ 0# 1# - - open IsIdempotentProsemiring isIdempotentProsemiring public - - idempotentSemiring : IdempotentSemiring c ℓ₁ - idempotentSemiring = record { isIdempotentSemiring = isIdempotentSemiring } - --- Preordered KleeneAlgebra (proKleeneAlgebra) - -record ProKleeneAlgebra c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infix 8 _⋆ - infixl 7 _*_ - infixl 6 _+_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ - _≤_ : Rel Carrier ℓ₂ - _+_ : Op₂ Carrier - _*_ : Op₂ Carrier - _⋆ : Op₁ Carrier - 0# : Carrier - 1# : Carrier - isProKleeneAlgebra : IsProKleeneAlgebra _≈_ _≤_ _+_ _*_ _⋆ 0# 1# - - open IsProKleeneAlgebra isProKleeneAlgebra public - - kleeneAlgebra : KleeneAlgebra c ℓ₁ - kleeneAlgebra = record { isKleeneAlgebra = isKleeneAlgebra } - ------------------------------------------------------------------------- --- Bundles of partially ordered structures - --- Partially ordered magmas (pomagmas) - -record Pomagma c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The partial order. - _∙_ : Op₂ Carrier -- Multiplication. - isPomagma : IsPomagma _≈_ _≤_ _∙_ - - open IsPomagma isPomagma public - - poset : Poset c ℓ₁ ℓ₂ - poset = record { isPartialOrder = isPartialOrder } - - promagma : Promagma c ℓ₁ ℓ₂ - promagma = record { isPromagma = isPromagma } - - open Promagma promagma public using (preorder; magma) - --- Partially ordered semigroups (posemigroups) - -record Posemigroup c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The partial order. - _∙_ : Op₂ Carrier -- Multiplication. - isPosemigroup : IsPosemigroup _≈_ _≤_ _∙_ - - open IsPosemigroup isPosemigroup public - - pomagma : Pomagma c ℓ₁ ℓ₂ - pomagma = record { isPomagma = isPomagma } - - open Pomagma pomagma public using (preorder; poset; magma; promagma) - - prosemigroup : Prosemigroup c ℓ₁ ℓ₂ - prosemigroup = record { isProsemigroup = isProsemigroup } - - open Prosemigroup prosemigroup public using (semigroup) - --- Partially ordered monoids (pomonoids) - -record Pomonoid c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The partial order. - _∙_ : Op₂ Carrier -- The monoid multiplication. - ε : Carrier -- The monoid unit. - isPomonoid : IsPomonoid _≈_ _≤_ _∙_ ε - - open IsPomonoid isPomonoid public - - posemigroup : Posemigroup c ℓ₁ ℓ₂ - posemigroup = record { isPosemigroup = isPosemigroup } - - open Posemigroup posemigroup public using - ( preorder - ; poset - ; magma - ; promagma - ; pomagma - ; semigroup - ; prosemigroup - ) - - promonoid : Promonoid c ℓ₁ ℓ₂ - promonoid = record { isPromonoid = isPromonoid } - - open Promonoid promonoid public using (monoid) - --- Partially ordered commutative monoids (commutative pomonoids) - -record CommutativePomonoid c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _∙_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ -- The underlying equality. - _≤_ : Rel Carrier ℓ₂ -- The partial order. - _∙_ : Op₂ Carrier -- The monoid multiplication. - ε : Carrier -- The monoid unit. - isCommutativePomonoid : IsCommutativePomonoid _≈_ _≤_ _∙_ ε - - open IsCommutativePomonoid isCommutativePomonoid public - - pomonoid : Pomonoid c ℓ₁ ℓ₂ - pomonoid = record { isPomonoid = isPomonoid } - - open Pomonoid pomonoid public using - ( preorder - ; poset - ; magma - ; promagma - ; pomagma - ; semigroup - ; prosemigroup - ; posemigroup - ; monoid - ; promonoid - ) - - commutativePromonoid : CommutativePromonoid c ℓ₁ ℓ₂ - commutativePromonoid = - record { isCommutativePromonoid = isCommutativePromonoid } - - open CommutativePromonoid commutativePromonoid public - using (commutativeSemigroup; commutativeMonoid) - --- Partially ordered semirings (posemirings) - -record Posemiring c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _*_ - infixl 6 _+_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ - _≤_ : Rel Carrier ℓ₂ - _+_ : Op₂ Carrier - _*_ : Op₂ Carrier - 0# : Carrier - 1# : Carrier - isPosemiring : IsPosemiring _≈_ _≤_ _+_ _*_ 0# 1# - - open IsPosemiring isPosemiring public - - +-commutativePomonoid : CommutativePomonoid c ℓ₁ ℓ₂ - +-commutativePomonoid = record - { isCommutativePomonoid = +-isCommutativePomonoid } - - open CommutativePomonoid +-commutativePomonoid public - using (preorder) - renaming - ( magma to +-magma - ; promagma to +-promagma - ; pomagma to +-pomagma - ; semigroup to +-semigroup - ; prosemigroup to +-prosemigroup - ; posemigroup to +-posemigroup - ; monoid to +-monoid - ; promonoid to +-promonoid - ; pomonoid to +-pomonoid - ; commutativeSemigroup to +-commutativeSemigroup - ; commutativeMonoid to +-commutativeMonoid - ; commutativePromonoid to +-commutativePromonoid - ) - - *-pomonoid : Pomonoid c ℓ₁ ℓ₂ - *-pomonoid = record { isPomonoid = *-isPomonoid } - - open Pomonoid *-pomonoid public - using () - renaming - ( magma to *-magma - ; promagma to *-promagma - ; pomagma to *-pomagma - ; semigroup to *-semigroup - ; prosemigroup to *-prosemigroup - ; posemigroup to *-posemigroup - ; monoid to *-monoid - ; promonoid to *-promonoid - ) - - prosemiring : Prosemiring c ℓ₁ ℓ₂ - prosemiring = record { isProsemiring = isProsemiring } - - open Prosemiring prosemiring public using (semiring) - --- Partially ordered idempotentSemiring (IdempotentPosemiring) - -record IdempotentPosemiring c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infixl 7 _*_ - infixl 6 _+_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ - _≤_ : Rel Carrier ℓ₂ - _+_ : Op₂ Carrier - _*_ : Op₂ Carrier - 0# : Carrier - 1# : Carrier - isIdempotentPosemiring : IsIdempotentPosemiring _≈_ _≤_ _+_ _*_ 0# 1# - - open IsIdempotentPosemiring isIdempotentPosemiring public - - idempotentProsemiring : IdempotentProsemiring c ℓ₁ ℓ₂ - idempotentProsemiring = record { isIdempotentProsemiring = isIdempotentProsemiring } - - open IdempotentProsemiring idempotentProsemiring public using (idempotentSemiring; +-idem) - --- Partially ordered KleeneAlgebra (PoKleeneAlgebra) - -record PoKleeneAlgebra c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - infix 4 _≈_ _≤_ - infix 8 _⋆ - infixl 7 _*_ - infixl 6 _+_ - field - Carrier : Set c - _≈_ : Rel Carrier ℓ₁ - _≤_ : Rel Carrier ℓ₂ - _+_ : Op₂ Carrier - _*_ : Op₂ Carrier - _⋆ : Op₁ Carrier - 0# : Carrier - 1# : Carrier - isPoKleeneAlgebra : IsPoKleeneAlgebra _≈_ _≤_ _+_ _*_ _⋆ 0# 1# - - open IsPoKleeneAlgebra isPoKleeneAlgebra public - - proKleeneAlgebra : ProKleeneAlgebra c ℓ₁ ℓ₂ - proKleeneAlgebra = record { isProKleeneAlgebra = isProKleeneAlgebra } - - open ProKleeneAlgebra proKleeneAlgebra public using (starExpansive; starDestructive) diff --git a/src/Algebra/Ordered/Structures.agda b/src/Algebra/Ordered/Structures.agda deleted file mode 100644 index 638d8783f0..0000000000 --- a/src/Algebra/Ordered/Structures.agda +++ /dev/null @@ -1,391 +0,0 @@ ------------------------------------------------------------------------- --- The Agda standard library --- --- Ordered algebraic structures (not packed up with sets, orders, --- operations, etc.) ------------------------------------------------------------------------- - --- The contents of this module should be accessed via --- `Algebra.Ordered`. - -{-# OPTIONS --cubical-compatible --safe #-} - -open import Relation.Binary.Core using (Rel; _⇒_) - -module Algebra.Ordered.Structures - {a ℓ₁ ℓ₂} {A : Set a} -- The underlying set - (_≈_ : Rel A ℓ₁) -- The underlying equality relation - (_≤_ : Rel A ℓ₂) -- The order - where - -open import Algebra.Core -open import Algebra.Definitions _≈_ -open import Algebra.Structures _≈_ -open import Data.Product.Base using (proj₁; proj₂) -open import Function.Base using (flip) -open import Level using (_⊔_) -open import Relation.Binary.Definitions using (Transitive; Monotonic₁; Monotonic₂) -open import Relation.Binary.Structures using (IsPreorder; IsPartialOrder) -open import Relation.Binary.Consequences using (mono₂⇒cong₂) - ------------------------------------------------------------------------- --- Preordered structures - --- Preordered magmas (promagmas) - -record IsPromagma (∙ : Op₂ A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPreorder : IsPreorder _≈_ _≤_ - ∙-cong : Congruent₂ ∙ - mono : Monotonic₂ _≤_ _≤_ _≤_ ∙ - - open IsPreorder isPreorder public - - mono₁ : ∀ {x} → Monotonic₁ _≤_ _≤_ (flip ∙ x) - mono₁ y≈z = mono y≈z refl - - mono₂ : ∀ {x} → Monotonic₁ _≤_ _≤_ (∙ x) - mono₂ y≈z = mono refl y≈z - - isMagma : IsMagma ∙ - isMagma = record { isEquivalence = isEquivalence ; ∙-cong = ∙-cong } - - open IsMagma isMagma public using (setoid; ∙-congˡ; ∙-congʳ) - --- Preordered semigroups (prosemigroups) - -record IsProsemigroup (∙ : Op₂ A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPromagma : IsPromagma ∙ - assoc : Associative ∙ - - open IsPromagma isPromagma public - - isSemigroup : IsSemigroup ∙ - isSemigroup = record { isMagma = isMagma ; assoc = assoc } - --- Preordered monoids (promonoids) - -record IsPromonoid (∙ : Op₂ A) (ε : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isProsemigroup : IsProsemigroup ∙ - identity : Identity ε ∙ - - open IsProsemigroup isProsemigroup public - - isMonoid : IsMonoid ∙ ε - isMonoid = record { isSemigroup = isSemigroup ; identity = identity } - - open IsMonoid isMonoid public using (identityˡ; identityʳ) - --- Preordered commutative monoids (commutative promonoids) - -record IsCommutativePromonoid (∙ : Op₂ A) (ε : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPromonoid : IsPromonoid ∙ ε - comm : Commutative ∙ - - open IsPromonoid isPromonoid public - - isCommutativeMonoid : IsCommutativeMonoid ∙ ε - isCommutativeMonoid = record { isMonoid = isMonoid ; comm = comm } - - open IsCommutativeMonoid isCommutativeMonoid public - using (isCommutativeSemigroup) - --- Preordered semirings (prosemirings) - -record IsProsemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - +-isCommutativePromonoid : IsCommutativePromonoid + 0# - *-cong : Congruent₂ * - *-mono : Monotonic₂ _≤_ _≤_ _≤_ * - *-assoc : Associative * - *-identity : Identity 1# * - distrib : * DistributesOver + - zero : Zero 0# * - - open IsCommutativePromonoid +-isCommutativePromonoid public - renaming - ( assoc to +-assoc - ; ∙-cong to +-cong - ; ∙-congˡ to +-congˡ - ; ∙-congʳ to +-congʳ - ; mono to +-mono - ; mono₁ to +-mono₁ - ; mono₂ to +-mono₂ - ; identity to +-identity - ; identityˡ to +-identityˡ - ; identityʳ to +-identityʳ - ; comm to +-comm - ; isMagma to +-isMagma - ; isSemigroup to +-isSemigroup - ; isMonoid to +-isMonoid - ; isCommutativeSemigroup to +-isCommutativeSemigroup - ; isCommutativeMonoid to +-isCommutativeMonoid - ) - - *-isPromonoid : IsPromonoid * 1# - *-isPromonoid = record - { isProsemigroup = record - { isPromagma = record - { isPreorder = isPreorder - ; ∙-cong = *-cong - ; mono = *-mono - } - ; assoc = *-assoc - } - ; identity = *-identity - } - - open IsPromonoid *-isPromonoid public - using () - renaming - ( ∙-congˡ to *-congˡ - ; ∙-congʳ to *-congʳ - ; mono₁ to *-mono₁ - ; mono₂ to *-mono₂ - ; identityˡ to *-identityˡ - ; identityʳ to *-identityʳ - ; isMagma to *-isMagma - ; isSemigroup to *-isSemigroup - ; isMonoid to *-isMonoid - ) - - isSemiring : IsSemiring + * 0# 1# - isSemiring = record - { isSemiringWithoutAnnihilatingZero = record - { +-isCommutativeMonoid = +-isCommutativeMonoid - ; *-cong = *-cong - ; *-assoc = *-assoc - ; *-identity = *-identity - ; distrib = distrib - } - ; zero = zero - } - - open IsSemiring isSemiring public using (distribˡ; distribʳ; zeroˡ; zeroʳ) - --- Preordered IdempotentSemiring (IdempotentProsemiring) - -record IsIdempotentProsemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isProsemiring : IsProsemiring + * 0# 1# - +-idem : Idempotent + - - open IsProsemiring isProsemiring public - - isIdempotentSemiring : IsIdempotentSemiring + * 0# 1# - isIdempotentSemiring = record { isSemiring = isSemiring ; +-idem = +-idem } - - open IsIdempotentSemiring isIdempotentSemiring public using (+-idem) - --- Preordered KleeneAlgebra (proKleeneAlgebra) - -record IsProKleeneAlgebra (+ * : Op₂ A) (⋆ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isIdempotentProsemiring : IsIdempotentProsemiring + * 0# 1# - starExpansive : StarExpansive 1# + * ⋆ - starDestructive : StarDestructive + * ⋆ - - open IsIdempotentProsemiring isIdempotentProsemiring public - - isKleeneAlgebra : IsKleeneAlgebra + * ⋆ 0# 1# - isKleeneAlgebra = record - { isIdempotentSemiring = isIdempotentSemiring - ; starExpansive = starExpansive - ; starDestructive = starDestructive - } - - open IsKleeneAlgebra isKleeneAlgebra public using (starExpansive; starDestructive) - ------------------------------------------------------------------------- --- Partially ordered structures - --- Partially ordered magmas (pomagmas) - -record IsPomagma (∙ : Op₂ A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPartialOrder : IsPartialOrder _≈_ _≤_ - mono : Monotonic₂ _≤_ _≤_ _≤_ ∙ - - open IsPartialOrder isPartialOrder public - - ∙-cong : Congruent₂ ∙ - ∙-cong = mono₂⇒cong₂ _≈_ _≈_ Eq.sym reflexive antisym mono - - isPromagma : IsPromagma ∙ - isPromagma = record - { isPreorder = isPreorder - ; ∙-cong = ∙-cong - ; mono = mono - } - - open IsPromagma isPromagma public - using (setoid; ∙-congˡ; ∙-congʳ; mono₁; mono₂; isMagma) - --- Partially ordered semigroups (posemigroups) - -record IsPosemigroup (∙ : Op₂ A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPomagma : IsPomagma ∙ - assoc : Associative ∙ - - open IsPomagma isPomagma public - - isProsemigroup : IsProsemigroup ∙ - isProsemigroup = record { isPromagma = isPromagma ; assoc = assoc } - - open IsProsemigroup isProsemigroup public using (isSemigroup) - --- Partially ordered monoids (pomonoids) - -record IsPomonoid (∙ : Op₂ A) (ε : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPosemigroup : IsPosemigroup ∙ - identity : Identity ε ∙ - - open IsPosemigroup isPosemigroup public - - isPromonoid : IsPromonoid ∙ ε - isPromonoid = record - { isProsemigroup = isProsemigroup - ; identity = identity - } - - open IsPromonoid isPromonoid public - using (isMonoid; identityˡ; identityʳ) - --- Partially ordered commutative monoids (commutative pomonoids) - -record IsCommutativePomonoid (∙ : Op₂ A) (ε : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPomonoid : IsPomonoid ∙ ε - comm : Commutative ∙ - - open IsPomonoid isPomonoid public - - isCommutativePromonoid : IsCommutativePromonoid ∙ ε - isCommutativePromonoid = record { isPromonoid = isPromonoid ; comm = comm } - - open IsCommutativePromonoid isCommutativePromonoid public - using (isCommutativeMonoid; isCommutativeSemigroup) - --- Partially ordered semirings (posemirings) - -record IsPosemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - +-isCommutativePomonoid : IsCommutativePomonoid + 0# - *-mono : Monotonic₂ _≤_ _≤_ _≤_ * - *-assoc : Associative * - *-identity : Identity 1# * - distrib : * DistributesOver + - zero : Zero 0# * - - open IsCommutativePomonoid +-isCommutativePomonoid public - renaming - ( assoc to +-assoc - ; ∙-cong to +-cong - ; ∙-congˡ to +-congˡ - ; ∙-congʳ to +-congʳ - ; mono to +-mono - ; mono₁ to +-mono₁ - ; mono₂ to +-mono₂ - ; identity to +-identity - ; identityˡ to +-identityˡ - ; identityʳ to +-identityʳ - ; comm to +-comm - ; isMagma to +-isMagma - ; isPromagma to +-isPromagma - ; isPomagma to +-isPomagma - ; isSemigroup to +-isSemigroup - ; isProsemigroup to +-isProsemigroup - ; isPosemigroup to +-isPosemigroup - ; isMonoid to +-isMonoid - ; isPromonoid to +-isPromonoid - ; isPomonoid to +-isPomonoid - ; isCommutativeSemigroup to +-isCommutativeSemigroup - ; isCommutativeMonoid to +-isCommutativeMonoid - ; isCommutativePromonoid to +-isCommutativePromonoid - ) - - *-isPomonoid : IsPomonoid * 1# - *-isPomonoid = record - { isPosemigroup = record - { isPomagma = record - { isPartialOrder = isPartialOrder - ; mono = *-mono - } - ; assoc = *-assoc - } - ; identity = *-identity - } - - open IsPomonoid *-isPomonoid public - using () - renaming - ( ∙-cong to *-cong - ; ∙-congˡ to *-congˡ - ; ∙-congʳ to *-congʳ - ; mono₁ to *-mono₁ - ; mono₂ to *-mono₂ - ; identityˡ to *-identityˡ - ; identityʳ to *-identityʳ - ; isMagma to *-isMagma - ; isPromagma to *-isPromagma - ; isPomagma to *-isPomagma - ; isSemigroup to *-isSemigroup - ; isProsemigroup to *-isProsemigroup - ; isPosemigroup to *-isPosemigroup - ; isMonoid to *-isMonoid - ; isPromonoid to *-isPromonoid - ) - - isProsemiring : IsProsemiring + * 0# 1# - isProsemiring = record - { +-isCommutativePromonoid = +-isCommutativePromonoid - ; *-cong = *-cong - ; *-mono = *-mono - ; *-assoc = *-assoc - ; *-identity = *-identity - ; distrib = distrib - ; zero = zero - } - - open IsProsemiring isProsemiring public - using (isSemiring; distribˡ; distribʳ; zeroˡ; zeroʳ) - --- Partially ordered idempotentSemiring (IdempotentPosemiring) - -record IsIdempotentPosemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isPosemiring : IsPosemiring + * 0# 1# - +-idem : Idempotent + - - open IsPosemiring isPosemiring public - - isIdempotentProsemiring : IsIdempotentProsemiring + * 0# 1# - isIdempotentProsemiring = record { isProsemiring = isProsemiring ; +-idem = +-idem } - - open IsIdempotentProsemiring isIdempotentProsemiring public - using (isIdempotentSemiring; +-idem) - --- Partially ordered KleeneAlgebra (PoKleeneAlgebra) - -record IsPoKleeneAlgebra (+ * : Op₂ A) (⋆ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - field - isIdempotentPosemiring : IsIdempotentPosemiring + * 0# 1# - starExpansive : StarExpansive 1# + * ⋆ - starDestructive : StarDestructive + * ⋆ - - open IsIdempotentPosemiring isIdempotentPosemiring public - - isProKleeneAlgebra : IsProKleeneAlgebra + * ⋆ 0# 1# - isProKleeneAlgebra = record - { isIdempotentProsemiring = isIdempotentProsemiring - ; starExpansive = starExpansive - ; starDestructive = starDestructive - } - - open IsProKleeneAlgebra isProKleeneAlgebra public - using (isKleeneAlgebra; starExpansive; starDestructive) diff --git a/src/Relation/Binary/Lattice/Properties/BoundedJoinSemilattice.agda b/src/Relation/Binary/Lattice/Properties/BoundedJoinSemilattice.agda index 276e3a9705..4d089b6b0f 100644 --- a/src/Relation/Binary/Lattice/Properties/BoundedJoinSemilattice.agda +++ b/src/Relation/Binary/Lattice/Properties/BoundedJoinSemilattice.agda @@ -14,13 +14,11 @@ module Relation.Binary.Lattice.Properties.BoundedJoinSemilattice open BoundedJoinSemilattice J open import Algebra.Definitions _≈_ -open import Algebra.Ordered.Structures using (IsCommutativePomonoid) -open import Algebra.Ordered.Bundles using (CommutativePomonoid) open import Data.Product.Base using (_,_) open import Function.Base using (_∘_; flip) open import Relation.Binary.Properties.Poset poset open import Relation.Binary.Lattice.Properties.JoinSemilattice joinSemilattice - using (isPosemigroup; ∨-comm) + using (∨-comm) -- Bottom is an identity of the meet operation. @@ -54,24 +52,3 @@ dualBoundedMeetSemilattice = record { ⊤ = ⊥ ; isBoundedMeetSemilattice = dualIsBoundedMeetSemilattice } - --- Every bounded semilattice gives rise to a commutative pomonoid - -isCommutativePomonoid : IsCommutativePomonoid _≈_ _≤_ _∨_ ⊥ -isCommutativePomonoid = record - { isPomonoid = record - { isPosemigroup = isPosemigroup - ; identity = identity - } - ; comm = ∨-comm - } - -commutativePomonoid : CommutativePomonoid c ℓ₁ ℓ₂ -commutativePomonoid = record - { Carrier = Carrier - ; _≈_ = _≈_ - ; _≤_ = _≤_ - ; _∙_ = _∨_ - ; ε = ⊥ - ; isCommutativePomonoid = isCommutativePomonoid - } diff --git a/src/Relation/Binary/Lattice/Properties/JoinSemilattice.agda b/src/Relation/Binary/Lattice/Properties/JoinSemilattice.agda index c68467f577..c94c3bad16 100644 --- a/src/Relation/Binary/Lattice/Properties/JoinSemilattice.agda +++ b/src/Relation/Binary/Lattice/Properties/JoinSemilattice.agda @@ -16,8 +16,6 @@ open JoinSemilattice J import Algebra.Lattice as Alg import Algebra.Structures as Alg open import Algebra.Definitions _≈_ -open import Algebra.Ordered.Structures using (IsPosemigroup) -open import Algebra.Ordered.Bundles using (Posemigroup) open import Data.Product.Base using (_,_) open import Function.Base using (_∘_; flip) open import Relation.Binary.Core using (_Preserves₂_⟶_⟶_) @@ -97,26 +95,6 @@ isAlgSemilattice = record algSemilattice : Alg.Semilattice c ℓ₁ algSemilattice = record { isSemilattice = isAlgSemilattice } --- Every semilattice gives rise to a posemigroup - -isPosemigroup : IsPosemigroup _≈_ _≤_ _∨_ -isPosemigroup = record - { isPomagma = record - { isPartialOrder = isPartialOrder - ; mono = ∨-monotonic - } - ; assoc = ∨-assoc - } - -posemigroup : Posemigroup c ℓ₁ ℓ₂ -posemigroup = record - { Carrier = Carrier - ; _≈_ = _≈_ - ; _≤_ = _≤_ - ; _∙_ = _∨_ - ; isPosemigroup = isPosemigroup - } - ------------------------------------------------------------------------ -- The dual construction is a meet semilattice. diff --git a/src/Relation/Binary/Lattice/Properties/MeetSemilattice.agda b/src/Relation/Binary/Lattice/Properties/MeetSemilattice.agda index a78d4fea6a..f8fca492b7 100644 --- a/src/Relation/Binary/Lattice/Properties/MeetSemilattice.agda +++ b/src/Relation/Binary/Lattice/Properties/MeetSemilattice.agda @@ -35,7 +35,7 @@ dualJoinSemilattice = record } open J dualJoinSemilattice public - using (isAlgSemilattice; algSemilattice; isPosemigroup; posemigroup) + using (isAlgSemilattice; algSemilattice) renaming ( ∨-monotonic to ∧-monotonic ; ∨-cong to ∧-cong From 135951fcde972e334150105fd9d54367bbe4b407 Mon Sep 17 00:00:00 2001 From: "G. Allais" Date: Fri, 13 Oct 2023 00:39:15 +0100 Subject: [PATCH 21/57] [ fix ] missing name in LICENCE file (#2139) --- LICENCE | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/LICENCE b/LICENCE index 75026ebb47..df31f24985 100644 --- a/LICENCE +++ b/LICENCE @@ -7,7 +7,8 @@ Eric Mertens, Joachim Breitner, Liyang Hu, Noam Zeilberger, Érdi Gergő, Stevan Andjelkovic, Helmut Grohne, Guilhem Moulin, Noriyuki Ohkawa, Evgeny Kotelnikov, James Chapman, Wen Kokke, Matthew Daggitt, Jason Hu, Sandro Stucki, Milo Turner, Zack Grannan, Lex van der Stoep, -Jacques Carette, James McKinna and some anonymous contributors. +Jacques Carette, James McKinna, Guillaume Allais +and some anonymous contributors. Permission is hereby granted, free of charge, to any person obtaining a copy of this software and associated documentation files (the From a843d0d4e1115d3ec16860f8d20022fe3aeefb90 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Am=C3=A9lia?= Date: Fri, 13 Oct 2023 00:41:38 -0300 Subject: [PATCH 22/57] Add new blocking primitives to `Reflection.TCM` (#1972) --- CHANGELOG.md | 9 +++++++++ src/Reflection/TCM.agda | 7 +++++-- 2 files changed, 14 insertions(+), 2 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 7755753af2..7938f8e94e 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -3812,6 +3812,15 @@ This is a full list of proofs that have changed form to use irrelevant instance WellFounded _<_ → Irreflexive _≈_ _<_ ``` +* Added new types and operations to `Reflection.TCM`: + ``` + Blocker : Set + blockerMeta : Meta → Blocker + blockerAny : List Blocker → Blocker + blockerAll : List Blocker → Blocker + blockTC : Blocker → TC A + ``` + * Added new file `Relation.Binary.Reasoning.Base.Apartness` This is how to use it: diff --git a/src/Reflection/TCM.agda b/src/Reflection/TCM.agda index a3df1f1bc3..d762dab2c5 100644 --- a/src/Reflection/TCM.agda +++ b/src/Reflection/TCM.agda @@ -8,8 +8,9 @@ module Reflection.TCM where -open import Reflection.AST.Term import Agda.Builtin.Reflection as Builtin + +open import Reflection.AST.Term import Reflection.TCM.Format as Format ------------------------------------------------------------------------ @@ -29,7 +30,9 @@ open Builtin public ; getContext; extendContext; inContext; freshName ; declareDef; declarePostulate; defineFun; getType; getDefinition ; blockOnMeta; commitTC; isMacro; withNormalisation - ; debugPrint; noConstraints; runSpeculative) + ; debugPrint; noConstraints; runSpeculative + ; Blocker; blockerMeta; blockerAny; blockerAll; blockTC + ) renaming (returnTC to pure) open Format public From aa6ef23369cbcf52d76c7f7dcfb14e5045cea48d Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Fri, 13 Oct 2023 18:22:01 +0900 Subject: [PATCH 23/57] Change definition of `IsStrictTotalOrder` (#2137) --- CHANGELOG.md | 40 ++++++++++++++ src/Data/Bool/Properties.agda | 5 +- src/Data/Char/Properties.agda | 5 +- src/Data/Fin/Properties.agda | 5 +- src/Data/Integer/Properties.agda | 5 +- src/Data/List/Relation/Binary/Lex/Strict.agda | 5 +- src/Data/Nat/Binary/Properties.agda | 5 +- src/Data/Nat/Properties.agda | 9 +--- .../Product/Relation/Binary/Lex/Strict.agda | 9 ++-- src/Data/Rational/Properties.agda | 5 +- .../Rational/Unnormalised/Properties.agda | 5 +- src/Data/Sum/Relation/Binary/LeftOrder.agda | 5 +- src/Data/Vec/Relation/Binary/Lex/Strict.agda | 3 +- src/Induction/WellFounded.agda | 4 +- src/Relation/Binary.agda | 1 + src/Relation/Binary/Bundles.agda | 2 +- .../Binary/Construct/Add/Infimum/Strict.agda | 10 ++-- .../Binary/Construct/Add/Supremum/Strict.agda | 10 ++-- .../Binary/Construct/Flip/EqAndOrd.agda | 5 +- src/Relation/Binary/Construct/Flip/Ord.agda | 5 +- .../Binary/Construct/NonStrictToStrict.agda | 5 +- src/Relation/Binary/Construct/On.agda | 5 +- .../Binary/Morphism/OrderMonomorphism.agda | 5 +- src/Relation/Binary/Structures.agda | 53 ++++++++++--------- src/Relation/Binary/Structures/Biased.agda | 49 +++++++++++++++++ 25 files changed, 164 insertions(+), 96 deletions(-) create mode 100644 src/Relation/Binary/Structures/Biased.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index 7938f8e94e..878a4fee2d 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -918,6 +918,46 @@ Non-backwards compatible changes * The old names (and the names of all proofs about these functions) have been deprecated appropriately. +### Changes to definition of `IsStrictTotalOrder` + +* The previous definition of the record `IsStrictTotalOrder` did not build upon `IsStrictPartialOrder` + as would be expected. Instead it omitted several fields like irreflexivity as they were derivable from the + proof of trichotomy. However, this led to problems further up the hierarchy where bundles such as `StrictTotalOrder` + which contained multiple distinct proofs of `IsStrictPartialOrder`. + +* To remedy this the definition of `IsStrictTotalOrder` has been changed to so that it builds upon + `IsStrictPartialOrder` as would be expected. + +* To aid migration, the old record definition has been moved to `Relation.Binary.Structures.Biased` + which contains the `isStrictTotalOrderᶜ` smart constructor (which is re-exported by `Relation.Binary`) . + Therefore the old code: + ```agda + <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ + <-isStrictTotalOrder = record + { isEquivalence = isEquivalence + ; trans = <-trans + ; compare = <-cmp + } + ``` + can be migrated either by updating to the new record fields if you have a proof of `IsStrictPartialOrder` + available: + ```agda + <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ + <-isStrictTotalOrder = record + { isStrictPartialOrder = <-isStrictPartialOrder + ; compare = <-cmp + } + ``` + or simply applying the smart constructor to the record definition as follows: + ```agda + <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ + <-isStrictTotalOrder = isStrictTotalOrderᶜ record + { isEquivalence = isEquivalence + ; trans = <-trans + ; compare = <-cmp + } + ``` + ### Changes to triple reasoning interface * The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof that the strict diff --git a/src/Data/Bool/Properties.agda b/src/Data/Bool/Properties.agda index 3c1c25d245..e911390366 100644 --- a/src/Data/Bool/Properties.agda +++ b/src/Data/Bool/Properties.agda @@ -210,9 +210,8 @@ true ; tri<; _Respects₂_) +open import Relation.Binary open import Relation.Binary.Consequences using (flip-Connex) open import Relation.Binary.PropositionalEquality open import Relation.Nullary hiding (Irrelevant) @@ -402,7 +397,7 @@ _>?_ = flip _?_ = flip _?_ = flip _ Date: Sat, 14 Oct 2023 09:30:28 +0900 Subject: [PATCH 24/57] Add _<$>_ operator for Function bundle (#2144) --- CHANGELOG.md | 6 ++++++ src/Function/Bundles.agda | 13 +++++++++++++ 2 files changed, 19 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index 878a4fee2d..8cd28d8f27 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -3159,6 +3159,12 @@ Additions to existing modules execState : State s a → s → s ``` +* Added new application operator synonym to `Function.Bundles`: + ```agda + _⟨$⟩_ : Func From To → Carrier From → Carrier To + _⟨$⟩_ = Func.to + ``` + * Added new proofs in `Function.Construct.Symmetry`: ``` bijective : Bijective ≈₁ ≈₂ f → Symmetric ≈₂ → Transitive ≈₂ → Congruent ≈₁ ≈₂ f → Bijective ≈₂ ≈₁ f⁻¹ diff --git a/src/Function/Bundles.agda b/src/Function/Bundles.agda index 1e479cf46a..a51dc4ea84 100644 --- a/src/Function/Bundles.agda +++ b/src/Function/Bundles.agda @@ -66,6 +66,7 @@ module _ (From : Setoid a ℓ₁) (To : Setoid b ℓ₂) where open IsCongruent isCongruent public using (module Eq₁; module Eq₂) + record Injection : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) where field to : A → B @@ -472,3 +473,15 @@ module _ {A : Set a} {B : Set b} where , strictlyInverseʳ⇒inverseʳ to invʳ ) +------------------------------------------------------------------------ +-- Other +------------------------------------------------------------------------ + +-- Alternative syntax for the application of functions + +module _ {From : Setoid a ℓ₁} {To : Setoid b ℓ₂} where + open Setoid + + infixl 5 _⟨$⟩_ + _⟨$⟩_ : Func From To → Carrier From → Carrier To + _⟨$⟩_ = Func.to From 222c2387382ece0f5a8dce69107c50913788e6e2 Mon Sep 17 00:00:00 2001 From: "G. Allais" Date: Sat, 14 Oct 2023 08:48:13 +0100 Subject: [PATCH 25/57] [ fix #2086 ] new web deployment strategy (#2147) --- .github/workflows/ci-ubuntu.yml | 6 +++++- travis/agda-logo.svg | 16 ++++++++++++++++ travis/landing-bottom.html | 4 ++++ travis/landing-top.html | 24 ++++++++++++++++++++++++ travis/landing.sh | 17 +++++++++++++++++ 5 files changed, 66 insertions(+), 1 deletion(-) create mode 100644 travis/agda-logo.svg create mode 100644 travis/landing-bottom.html create mode 100644 travis/landing-top.html create mode 100755 travis/landing.sh diff --git a/.github/workflows/ci-ubuntu.yml b/.github/workflows/ci-ubuntu.yml index ebafe70fa8..2ecc969589 100644 --- a/.github/workflows/ci-ubuntu.yml +++ b/.github/workflows/ci-ubuntu.yml @@ -75,7 +75,7 @@ jobs: || '${{ github.base_ref }}' == 'master' ]]; then # Pick Agda version for master echo "AGDA_COMMIT=tags/v2.6.4" >> $GITHUB_ENV; - echo "AGDA_HTML_DIR=html" >> $GITHUB_ENV + echo "AGDA_HTML_DIR=html/master" >> $GITHUB_ENV elif [[ '${{ github.ref }}' == 'refs/heads/experimental' \ || '${{ github.base_ref }}' == 'experimental' ]]; then # Pick Agda version for experimental @@ -125,6 +125,7 @@ jobs: run: cabal update - name: Install alex & happy + if: steps.cache-cabal.outputs.cache-hit != 'true' run: | ${{ env.CABAL_INSTALL }} alex ${{ env.CABAL_INSTALL }} happy @@ -177,6 +178,9 @@ jobs: rm -f '${{ env.AGDA_HTML_DIR }}'/*.css ${{ env.AGDA }} --html --html-dir ${{ env.AGDA_HTML_DIR }} index.agda + cp travis/* . + ./landing.sh + - name: Deploy HTML uses: JamesIves/github-pages-deploy-action@4.1.3 if: ${{ success() && env.AGDA_DEPLOY }} diff --git a/travis/agda-logo.svg b/travis/agda-logo.svg new file mode 100644 index 0000000000..0fe434e8ed --- /dev/null +++ b/travis/agda-logo.svg @@ -0,0 +1,16 @@ + + logotype + Created with Sketch. + + + + \ No newline at end of file diff --git a/travis/landing-bottom.html b/travis/landing-bottom.html new file mode 100644 index 0000000000..9cd52d7c69 --- /dev/null +++ b/travis/landing-bottom.html @@ -0,0 +1,4 @@ + + + + diff --git a/travis/landing-top.html b/travis/landing-top.html new file mode 100644 index 0000000000..81b03304b9 --- /dev/null +++ b/travis/landing-top.html @@ -0,0 +1,24 @@ + + + + Documention for the Agda standard library + + + + +
+ +

Documention for the Agda standard library

+ +
+ +

Development versions

+ + + +

Released versions

+ +
    diff --git a/travis/landing.sh b/travis/landing.sh new file mode 100755 index 0000000000..55ebe0c8fb --- /dev/null +++ b/travis/landing.sh @@ -0,0 +1,17 @@ +set -eu +set -o pipefail + +rm html/index.html + +cat landing-top.html >> landing.html + +find html/ -name "index.html" \ + | grep -v "master\|experimental" \ + | sort -r \ + | sed 's|html/\([^\/]*\)/index.html|
  • \1
  • |g' \ + >> landing.html + +cat landing-bottom.html >> landing.html + +mv landing.html html/index.html +mv agda-logo.svg html/ From 347a0798e240336811adeb022cb238053a3f7596 Mon Sep 17 00:00:00 2001 From: "G. Allais" Date: Sun, 15 Oct 2023 04:59:09 +0100 Subject: [PATCH 26/57] [ admin ] fix sorting logic (#2151) With the previous script we were sorting entries of the form html/vX.Y.Z/index.html but the order is such that vX.Y/ < vX.Y.Z/ and so we were ending up with v1.7 coming after v1.7.3. This fixes that by using sed to get rid of the html/ prefix and the /index.html suffix before the sorting phase. --- travis/landing.sh | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/travis/landing.sh b/travis/landing.sh index 55ebe0c8fb..fcc042da46 100755 --- a/travis/landing.sh +++ b/travis/landing.sh @@ -7,8 +7,9 @@ cat landing-top.html >> landing.html find html/ -name "index.html" \ | grep -v "master\|experimental" \ + | sed 's|html/\([^\/]*\)/index.html|\1|g' \ | sort -r \ - | sed 's|html/\([^\/]*\)/index.html|
  • \1
  • |g' \ + | sed 's|^\(.*\)$|
  • \1
  • |g' \ >> landing.html cat landing-bottom.html >> landing.html From d91e21d379734eb9230aaa306a1a24c6ee2aebe0 Mon Sep 17 00:00:00 2001 From: Nathan van Doorn Date: Mon, 16 Oct 2023 05:13:47 +0200 Subject: [PATCH 27/57] Add coincidence law to modules (#1898) --- CHANGELOG.md | 5 +++++ .../Module/Construct/DirectProduct.agda | 2 ++ src/Algebra/Module/Construct/TensorUnit.agda | 2 ++ src/Algebra/Module/Definitions.agda | 2 ++ .../Module/Definitions/Bi/Simultaneous.agda | 18 ++++++++++++++++++ src/Algebra/Module/Structures.agda | 18 +++++++++++------- src/Algebra/Module/Structures/Biased.agda | 12 ++++++++++-- 7 files changed, 50 insertions(+), 9 deletions(-) create mode 100644 src/Algebra/Module/Definitions/Bi/Simultaneous.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index 8cd28d8f27..c6479e6ce7 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1843,6 +1843,11 @@ New modules Algebra.Module.Core ``` +* Definitions for module-like algebraic structures with left- and right- multiplication over the same scalars: + ``` + Algebra.Module.Definitions.Bi.Simultaneous + ``` + * Constructive algebraic structures with apartness relations: ``` Algebra.Apartness diff --git a/src/Algebra/Module/Construct/DirectProduct.agda b/src/Algebra/Module/Construct/DirectProduct.agda index 840cc42cab..429a2d11ca 100644 --- a/src/Algebra/Module/Construct/DirectProduct.agda +++ b/src/Algebra/Module/Construct/DirectProduct.agda @@ -97,6 +97,7 @@ semimodule M N = record { isSemimodule = record { isBisemimodule = Bisemimodule.isBisemimodule (bisemimodule M.bisemimodule N.bisemimodule) + ; *ₗ-*ᵣ-coincident = λ x m → M.*ₗ-*ᵣ-coincident x (proj₁ m) , N.*ₗ-*ᵣ-coincident x (proj₂ m) } } where module M = Semimodule M; module N = Semimodule N @@ -155,5 +156,6 @@ bimodule M N = record ⟨module⟩ M N = record { isModule = record { isBimodule = Bimodule.isBimodule (bimodule M.bimodule N.bimodule) + ; *ₗ-*ᵣ-coincident = λ x m → M.*ₗ-*ᵣ-coincident x (proj₁ m) , N.*ₗ-*ᵣ-coincident x (proj₂ m) } } where module M = Module M; module N = Module N diff --git a/src/Algebra/Module/Construct/TensorUnit.agda b/src/Algebra/Module/Construct/TensorUnit.agda index 6277210520..c76f2a39ba 100644 --- a/src/Algebra/Module/Construct/TensorUnit.agda +++ b/src/Algebra/Module/Construct/TensorUnit.agda @@ -77,6 +77,7 @@ semimodule : {R : CommutativeSemiring c ℓ} → Semimodule R c ℓ semimodule {R = commutativeSemiring} = record { isSemimodule = record { isBisemimodule = Bisemimodule.isBisemimodule bisemimodule + ; *ₗ-*ᵣ-coincident = *-comm } } where open CommutativeSemiring commutativeSemiring @@ -113,5 +114,6 @@ bimodule {R = ring} = record ⟨module⟩ {R = commutativeRing} = record { isModule = record { isBimodule = Bimodule.isBimodule bimodule + ; *ₗ-*ᵣ-coincident = *-comm } } where open CommutativeRing commutativeRing diff --git a/src/Algebra/Module/Definitions.agda b/src/Algebra/Module/Definitions.agda index 19c1abdb46..949aa45119 100644 --- a/src/Algebra/Module/Definitions.agda +++ b/src/Algebra/Module/Definitions.agda @@ -12,7 +12,9 @@ module Algebra.Module.Definitions where import Algebra.Module.Definitions.Left as L import Algebra.Module.Definitions.Right as R import Algebra.Module.Definitions.Bi as B + import Algebra.Module.Definitions.Bi.Simultaneous as BS module LeftDefs = L module RightDefs = R module BiDefs = B + module SimultaneousBiDefs = BS diff --git a/src/Algebra/Module/Definitions/Bi/Simultaneous.agda b/src/Algebra/Module/Definitions/Bi/Simultaneous.agda new file mode 100644 index 0000000000..686f4ff06c --- /dev/null +++ b/src/Algebra/Module/Definitions/Bi/Simultaneous.agda @@ -0,0 +1,18 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Properties connecting left-scaling and right-scaling over the same scalars +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +open import Relation.Binary + +module Algebra.Module.Definitions.Bi.Simultaneous + {a b ℓb} (A : Set a) {B : Set b} (_≈_ : Rel B ℓb) + where + +open import Algebra.Module.Core + +Coincident : Opₗ A B → Opᵣ A B → Set _ +Coincident _∙ₗ_ _∙ᵣ_ = ∀ x m → (x ∙ₗ m) ≈ (m ∙ᵣ x) diff --git a/src/Algebra/Module/Structures.agda b/src/Algebra/Module/Structures.agda index d8a291bd13..64873252bd 100644 --- a/src/Algebra/Module/Structures.agda +++ b/src/Algebra/Module/Structures.agda @@ -166,14 +166,16 @@ module _ (commutativeSemiring : CommutativeSemiring r ℓr) open CommutativeSemiring commutativeSemiring renaming (Carrier to R) -- An R-semimodule is an R-R-bisemimodule where R is commutative. - -- This means that *ₗ and *ᵣ coincide up to mathematical equality, - -- though it may be that they do not coincide up to definitional - -- equality. + -- We enforce that *ₗ and *ᵣ coincide up to mathematical equality, though it + -- may be that they do not coincide up to definitional equality. + + open SimultaneousBiDefs R ≈ᴹ record IsSemimodule (*ₗ : Opₗ R M) (*ᵣ : Opᵣ R M) : Set (r ⊔ m ⊔ ℓr ⊔ ℓm) where field isBisemimodule : IsBisemimodule semiring semiring ≈ᴹ +ᴹ 0ᴹ *ₗ *ᵣ + *ₗ-*ᵣ-coincident : Coincident *ₗ *ᵣ open IsBisemimodule isBisemimodule public @@ -282,15 +284,17 @@ module _ (commutativeRing : CommutativeRing r ℓr) open CommutativeRing commutativeRing renaming (Carrier to R) -- An R-module is an R-R-bimodule where R is commutative. - -- This means that *ₗ and *ᵣ coincide up to mathematical equality, - -- though it may be that they do not coincide up to definitional - -- equality. + -- We enforce that *ₗ and *ᵣ coincide up to mathematical equality, though it + -- may be that they do not coincide up to definitional equality. + + open SimultaneousBiDefs R ≈ᴹ record IsModule (*ₗ : Opₗ R M) (*ᵣ : Opᵣ R M) : Set (r ⊔ m ⊔ ℓr ⊔ ℓm) where field isBimodule : IsBimodule ring ring ≈ᴹ +ᴹ 0ᴹ -ᴹ *ₗ *ᵣ + *ₗ-*ᵣ-coincident : Coincident *ₗ *ᵣ open IsBimodule isBimodule public isSemimodule : IsSemimodule commutativeSemiring ≈ᴹ +ᴹ 0ᴹ *ₗ *ᵣ - isSemimodule = record { isBisemimodule = isBisemimodule } + isSemimodule = record { isBisemimodule = isBisemimodule; *ₗ-*ᵣ-coincident = *ₗ-*ᵣ-coincident } diff --git a/src/Algebra/Module/Structures/Biased.agda b/src/Algebra/Module/Structures/Biased.agda index 99e48695ae..764745c54c 100644 --- a/src/Algebra/Module/Structures/Biased.agda +++ b/src/Algebra/Module/Structures/Biased.agda @@ -57,7 +57,10 @@ module _ (commutativeSemiring : CommutativeSemiring r ℓr) where } isSemimodule : IsSemimodule commutativeSemiring ≈ᴹ +ᴹ 0ᴹ *ₗ (flip *ₗ) - isSemimodule = record { isBisemimodule = isBisemimodule } + isSemimodule = record + { isBisemimodule = isBisemimodule + ; *ₗ-*ᵣ-coincident = λ _ _ → ≈ᴹ-refl + } -- Similarly, a right semimodule over a commutative semiring -- is already a semimodule. @@ -88,7 +91,10 @@ module _ (commutativeSemiring : CommutativeSemiring r ℓr) where } isSemimodule : IsSemimodule commutativeSemiring ≈ᴹ +ᴹ 0ᴹ (flip *ᵣ) *ᵣ - isSemimodule = record { isBisemimodule = isBisemimodule } + isSemimodule = record + { isBisemimodule = isBisemimodule + ; *ₗ-*ᵣ-coincident = λ _ _ → ≈ᴹ-refl + } module _ (commutativeRing : CommutativeRing r ℓr) where @@ -113,6 +119,7 @@ module _ (commutativeRing : CommutativeRing r ℓr) where ; -ᴹ‿cong = -ᴹ‿cong ; -ᴹ‿inverse = -ᴹ‿inverse } + ; *ₗ-*ᵣ-coincident = λ _ _ → ≈ᴹ-refl } -- Similarly, a right module over a commutative ring is already a module. @@ -134,4 +141,5 @@ module _ (commutativeRing : CommutativeRing r ℓr) where ; -ᴹ‿cong = -ᴹ‿cong ; -ᴹ‿inverse = -ᴹ‿inverse } + ; *ₗ-*ᵣ-coincident = λ _ _ → ≈ᴹ-refl } From 010498b69e094876c52858184493a93a7a618611 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Mon, 16 Oct 2023 20:28:06 +0900 Subject: [PATCH 28/57] Make reasoning modular by adding new `Reasoning.Syntax` module (#2152) --- CHANGELOG.md | 28 ++ .../Stream/Relation/Binary/Pointwise.agda | 4 + src/Codata/Musical/Colist.agda | 32 +- .../Relation/Unary/Any/Properties.agda | 19 +- src/Data/Integer/Divisibility/Signed.agda | 11 +- src/Data/Integer/Properties.agda | 2 +- .../Relation/Binary/BagAndSetEquality.agda | 29 +- .../Binary/Permutation/Propositional.agda | 16 +- .../Relation/Binary/Permutation/Setoid.agda | 26 +- .../Binary/Subset/Setoid/Properties.agda | 29 +- src/Data/Nat/Binary/Properties.agda | 13 +- src/Data/Nat/Divisibility.agda | 11 +- src/Data/Nat/Properties.agda | 2 +- src/Data/Rational/Base.agda | 7 +- src/Data/Rational/Properties.agda | 58 +-- .../Rational/Unnormalised/Properties.agda | 24 +- .../Vec/Relation/Binary/Lex/NonStrict.agda | 2 +- src/Data/Vec/Relation/Binary/Lex/Strict.agda | 28 +- src/Function/Related/Propositional.agda | 53 ++- .../ReflexiveTransitive/Properties.agda | 19 +- .../Binary/HeterogeneousEquality.agda | 48 +-- .../Binary/Reasoning/Base/Apartness.agda | 113 +++--- .../Binary/Reasoning/Base/Double.agda | 110 ++---- .../Binary/Reasoning/Base/Partial.agda | 84 ++-- .../Binary/Reasoning/Base/Single.agda | 83 ++-- .../Binary/Reasoning/Base/Triple.agda | 164 +++----- .../Binary/Reasoning/MultiSetoid.agda | 65 ++-- .../Binary/Reasoning/PartialOrder.agda | 2 +- .../Binary/Reasoning/PartialSetoid.agda | 27 +- src/Relation/Binary/Reasoning/Setoid.agda | 23 +- .../Binary/Reasoning/StrictPartialOrder.agda | 2 +- src/Relation/Binary/Reasoning/Syntax.agda | 366 ++++++++++++++++++ 32 files changed, 856 insertions(+), 644 deletions(-) create mode 100644 src/Relation/Binary/Reasoning/Syntax.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index c6479e6ce7..efce415622 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -33,6 +33,9 @@ Bug-fixes in `Function.Construct.Composition` had their arguments in the wrong order. They have been flipped so they can actually be used as a composition operator. +* The combinators `_≃⟨_⟩_` and `_≃˘⟨_⟩_` in `Data.Rational.Properties.≤-Reasoning` + now correctly accepts proofs of type `_≃_` instead of the previous proofs of type `_≡_`. + * The following operators were missing a fixity declaration, which has now been fixed - ``` @@ -536,6 +539,10 @@ Non-backwards compatible changes 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. +* Similarly, in order to prevent reduction, the equality `_≃_` in `Data.Rational.Base` + has been made into a data type with the single constructor `*≡*`. The destructor `drop-*≡*` + has been added to `Data.Rational.Properties`. + ### Change to the definition of `LeftCancellative` and `RightCancellative` * The definitions of the types for cancellativity in `Algebra.Definitions` previously @@ -1222,6 +1229,15 @@ Major improvements * In `Relation.Binary.Reasoning.Base.Triple`, added a new parameter `<-irrefl : Irreflexive _≈_ _<_` +### More modular design of equational reasoning. + +* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports + a range of modules containing pre-existing reasoning combinator syntax. + +* This makes it possible to add new or rename existing reasoning combinators to a + pre-existing `Reasoning` module in just a couple of lines + (e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`) + Deprecated modules ------------------ @@ -1793,6 +1809,11 @@ Deprecated names sym-↔ ↦ ↔-sym ``` +* In `Function.Related.Propositional.Reasoning`: + ```agda + _↔⟨⟩_ ↦ _≡⟨⟩_ + ``` + * In `Foreign.Haskell.Either` and `Foreign.Haskell.Pair`: ``` toForeign ↦ Foreign.Haskell.Coerce.coerce @@ -2690,6 +2711,8 @@ Additions to existing modules toℕ-inverseˡ : Inverseˡ _≡_ _≡_ toℕ fromℕ toℕ-inverseʳ : Inverseʳ _≡_ _≡_ toℕ fromℕ toℕ-inverseᵇ : Inverseᵇ _≡_ _≡_ toℕ fromℕ + + <-asym : Asymmetric _<_ ``` * Added a new pattern synonym to `Data.Nat.Divisibility.Core`: @@ -2859,6 +2882,7 @@ Additions to existing modules ```agda ↥ᵘ-toℚᵘ : ↥ᵘ (toℚᵘ p) ≡ ↥ p ↧ᵘ-toℚᵘ : ↧ᵘ (toℚᵘ p) ≡ ↧ p + ↥p≡↥q≡0⇒p≡q : ↥ p ≡ 0ℤ → ↥ q ≡ 0ℤ → p ≡ q _≥?_ : Decidable _≥_ _>?_ : Decidable _>_ @@ -2883,6 +2907,10 @@ Additions to existing modules * Added new definitions in `Data.Rational.Unnormalised.Properties`: ```agda + ↥p≡0⇒p≃0 : ↥ p ≡ 0ℤ → p ≃ 0ℚᵘ + p≃0⇒↥p≡0 : p ≃ 0ℚᵘ → ↥ p ≡ 0ℤ + ↥p≡↥q≡0⇒p≃q : ↥ p ≡ 0ℤ → ↥ q ≡ 0ℤ → p ≃ q + +-*-rawNearSemiring : RawNearSemiring 0ℓ 0ℓ +-*-rawSemiring : RawSemiring 0ℓ 0ℓ diff --git a/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda b/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda index acc7153f52..9f573cc711 100644 --- a/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda +++ b/src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda @@ -127,10 +127,14 @@ module _ {A : Set a} where ------------------------------------------------------------------------ -- Pointwise DSL +-- -- A guardedness check does not play well with compositional proofs. -- Luckily we can learn from Nils Anders Danielsson's -- Beating the Productivity Checker Using Embedded Languages -- and design a little compositional DSL to define such proofs +-- +-- NOTE: also because of the guardedness check we can't use the standard +-- `Relation.Binary.Reasoning.Syntax` approach. module pw-Reasoning (S : Setoid a ℓ) where private module S = Setoid S diff --git a/src/Codata/Musical/Colist.agda b/src/Codata/Musical/Colist.agda index 2f59082e22..aeff3076dc 100644 --- a/src/Codata/Musical/Colist.agda +++ b/src/Codata/Musical/Colist.agda @@ -34,6 +34,7 @@ import Relation.Binary.Construct.FromRel as Ind import Relation.Binary.Reasoning.Preorder as PreR import Relation.Binary.Reasoning.PartialOrder as POR open import Relation.Binary.PropositionalEquality as P using (_≡_) +open import Relation.Binary.Reasoning.Syntax open import Relation.Nullary.Reflects using (invert) open import Relation.Nullary open import Relation.Nullary.Negation @@ -165,12 +166,11 @@ Any-∈ {P = P} = mk↔ₛ′ module ⊑-Reasoning {a} {A : Set a} where private module Base = POR (⊑-Poset A) - open Base public - hiding (step-<; begin-strict_; step-≤) + open Base public hiding (step-<; step-≤) + + open ⊑-syntax _IsRelatedTo_ _IsRelatedTo_ Base.≤-go public + open ⊏-syntax _IsRelatedTo_ _IsRelatedTo_ Base.<-go public - infixr 2 step-⊑ - step-⊑ = Base.step-≤ - syntax step-⊑ x ys⊑zs xs⊑ys = x ⊑⟨ xs⊑ys ⟩ ys⊑zs -- The subset relation forms a preorder. @@ -179,22 +179,22 @@ module ⊑-Reasoning {a} {A : Set a} where (λ xs≈ys → ⊑⇒⊆ (⊑P.reflexive xs≈ys)) where module ⊑P = Poset (⊑-Poset A) +-- Example uses: +-- +-- x∈zs : x ∈ zs +-- x∈zs = +-- x ∈⟨ x∈xs ⟩ +-- xs ⊆⟨ xs⊆ys ⟩ +-- ys ≡⟨ ys≡zs ⟩ +-- zs ∎ module ⊆-Reasoning {A : Set a} where private module Base = PreR (⊆-Preorder A) - open Base public - hiding (step-∼) - - infixr 2 step-⊆ - infix 1 step-∈ - - step-⊆ = Base.step-∼ + open Base public hiding (step-∼) - step-∈ : ∀ (x : A) {xs ys} → xs IsRelatedTo ys → x ∈ xs → x ∈ ys - step-∈ x xs⊆ys x∈xs = (begin xs⊆ys) x∈xs + open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → begin x) + open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public - syntax step-⊆ xs ys⊆zs xs⊆ys = xs ⊆⟨ xs⊆ys ⟩ ys⊆zs - syntax step-∈ x xs⊆ys x∈xs = x ∈⟨ x∈xs ⟩ xs⊆ys -- take returns a prefix. take-⊑ : ∀ n (xs : Colist A) → diff --git a/src/Data/Container/Relation/Unary/Any/Properties.agda b/src/Data/Container/Relation/Unary/Any/Properties.agda index d3265acefb..0a4340fa7b 100644 --- a/src/Data/Container/Relation/Unary/Any/Properties.agda +++ b/src/Data/Container/Relation/Unary/Any/Properties.agda @@ -26,7 +26,6 @@ open import Relation.Binary.Core using (REL) open import Relation.Binary.PropositionalEquality as P using (_≡_; _≗_; refl) -open Related.EquationalReasoning hiding (_≡⟨_⟩_) private module ×⊎ {k ℓ} = CommutativeSemiring (×-⊎-commutativeSemiring k ℓ) @@ -71,6 +70,7 @@ module _ {s p} {C : Container s p} {x} {X : Set x} (∃ λ x → x ∈ xs₁ × P₁ x) ∼⟨ Σ.cong ↔-refl (xs₁≈xs₂ ×-cong P₁↔P₂ _) ⟩ (∃ λ x → x ∈ xs₂ × P₂ x) ↔⟨ SK-sym (↔∈ C) ⟩ ◇ C P₂ xs₂ ∎ + where open Related.EquationalReasoning -- Nested occurrences of ◇ can sometimes be swapped. @@ -95,6 +95,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s (∃ λ y → y ∈ ys × ∃ λ x → x ∈ xs × P x y) ↔⟨ Σ.cong ↔-refl (Σ.cong ↔-refl (SK-sym (↔∈ C₁))) ⟩ (∃ λ y → y ∈ ys × ◇ _ (flip P y) xs) ↔⟨ SK-sym (↔∈ C₂) ⟩ ◇ _ (λ y → ◇ _ (flip P y) xs) ys ∎ + where open Related.EquationalReasoning -- Nested occurrences of ◇ can sometimes be flattened. @@ -162,9 +163,10 @@ module _ {s p} (C : Container s p) {x y} {X : Set x} {Y : Set y} map↔∘ : ∀ {xs : ⟦ C ⟧ X} (f : X → Y) → ◇ C P (map f xs) ↔ ◇ C (P ∘′ f) xs map↔∘ {xs} f = ◇ C P (map f xs) ↔⟨ ↔Σ C ⟩ - ∃ (P ∘′ proj₂ (map f xs)) ↔⟨⟩ + ∃ (P ∘′ proj₂ (map f xs)) ≡⟨⟩ ∃ (P ∘′ f ∘′ proj₂ xs) ↔⟨ SK-sym (↔Σ C) ⟩ ◇ C (P ∘′ f) xs ∎ + where open Related.EquationalReasoning -- Membership in a mapped container can be expressed without reference -- to map. @@ -178,6 +180,7 @@ module _ {s p} (C : Container s p) {x y} {X : Set x} {Y : Set y} y ∈ map f xs ↔⟨ map↔∘ C (y ≡_) f ⟩ ◇ C (λ x → y ≡ f x) xs ↔⟨ ↔∈ C ⟩ ∃ (λ x → x ∈ xs × y ≡ f x) ∎ + where open Related.EquationalReasoning -- map is a congruence for bag and set equality and related preorders. @@ -193,6 +196,7 @@ module _ {s p} (C : Container s p) {x y} {X : Set x} {Y : Set y} ◇ C (λ y → x ≡ f₂ y) xs₂ ↔⟨ SK-sym (map↔∘ C (_≡_ x) f₂) ⟩ x ∈ map f₂ xs₂ ∎ where + open Related.EquationalReasoning helper : ∀ y → (x ≡ f₁ y) ↔ (x ≡ f₂ y) helper y rewrite f₁≗f₂ y = ↔-refl @@ -205,7 +209,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s remove-linear {xs} m = mk↔ₛ′ t f t∘f f∘t where open _≃_ - open P.≡-Reasoning renaming (_∎ to _∎′) + open P.≡-Reasoning position⊸m : ∀ {s} → Position C₂ (shape⊸ m s) ≃ Position C₁ s position⊸m = ↔⇒≃ (position⊸ m) @@ -259,7 +263,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s P.subst (P ∘ proj₂ xs) P.refl p ≡⟨⟩ - p ∎′) + p ∎) ) t∘f : t ∘ f ≗ id @@ -279,7 +283,7 @@ module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s (P.trans-symˡ (right-inverse-of position⊸m _)) ⟩ P.subst (P ∘ proj₂ xs) P.refl p ≡⟨⟩ - p ∎′) + p ∎) ) -- Linear endomorphisms are identity functions if bag equality is used. @@ -290,6 +294,7 @@ module _ {s p} {C : Container s p} {x} {X : Set x} where linear-identity {xs} m {x} = x ∈ ⟪ m ⟫⊸ xs ↔⟨ remove-linear (_≡_ x) m ⟩ x ∈ xs ∎ + where open Related.EquationalReasoning -- If join can be expressed using a linear morphism (in a certain -- way), then it can be absorbed by the predicate. @@ -307,4 +312,6 @@ module _ {s₁ s₂ s₃ p₁ p₂ p₃} ◇ C₃ P (⟪ join ⟫⊸ xss′) ↔⟨ remove-linear P join ⟩ ◇ (C₁ C.∘ C₂) P xss′ ↔⟨ SK-sym $ flatten P xss ⟩ ◇ C₁ (◇ C₂ P) xss ∎ - where xss′ = Inverse.from (Composition.correct C₁ C₂) xss + where + open Related.EquationalReasoning + xss′ = Inverse.from (Composition.correct C₁ C₂) xss diff --git a/src/Data/Integer/Divisibility/Signed.agda b/src/Data/Integer/Divisibility/Signed.agda index af06d57486..9a4b938950 100644 --- a/src/Data/Integer/Divisibility/Signed.agda +++ b/src/Data/Integer/Divisibility/Signed.agda @@ -30,6 +30,7 @@ open import Relation.Binary.PropositionalEquality import Relation.Binary.Reasoning.Preorder as PreorderReasoning open import Relation.Nullary.Decidable using (yes; no) import Relation.Nullary.Decidable as DEC +open import Relation.Binary.Reasoning.Syntax ------------------------------------------------------------------------ -- Type @@ -118,15 +119,11 @@ open _∣_ using (quotient) public -- Divisibility reasoning module ∣-Reasoning where - private - module Base = PreorderReasoning ∣-preorder + private module Base = PreorderReasoning ∣-preorder - open Base public - hiding (step-∼; step-≈; step-≈˘) + open Base public hiding (step-∼; step-≈; step-≈˘) - infixr 2 step-∣ - step-∣ = Base.step-∼ - syntax step-∣ x y∣z x∣y = x ∣⟨ x∣y ⟩ y∣z + open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public ------------------------------------------------------------------------ -- Other properties of _∣_ diff --git a/src/Data/Integer/Properties.agda b/src/Data/Integer/Properties.agda index bfe29c4079..13ed3deab8 100644 --- a/src/Data/Integer/Properties.agda +++ b/src/Data/Integer/Properties.agda @@ -364,7 +364,7 @@ i≮i = <-irrefl refl module ≤-Reasoning where open import Relation.Binary.Reasoning.Base.Triple ≤-isPreorder - <-irrefl + <-asym <-trans (resp₂ _<_) <⇒≤ diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 889efe96d9..0acef52dcb 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -37,11 +37,13 @@ open import Function.Related.TypeIsomorphisms open import Function.Properties.Inverse using (↔-sym; ↔-trans; to-from) open import Level using (Level) open import Relation.Binary.Core using (_⇒_) +open import Relation.Binary.Definitions using (Trans) open import Relation.Binary.Bundles using (Preorder; Setoid) import Relation.Binary.Reasoning.Setoid as EqR import Relation.Binary.Reasoning.Preorder as PreorderReasoning open import Relation.Binary.PropositionalEquality as P using (_≡_; _≢_; _≗_; refl) +open import Relation.Binary.Reasoning.Syntax open import Relation.Nullary open import Data.List.Membership.Propositional.Properties @@ -90,29 +92,22 @@ bag-=⇒ xs≈ys = ↔⇒ xs≈ys ------------------------------------------------------------------------ -- "Equational" reasoning for _⊆_ along with an additional relatedness -module ⊆-Reasoning where - private - module PreOrder {a} {A : Set a} = PreorderReasoning (⊆-preorder A) +module ⊆-Reasoning {A : Set a} where + private module Base = PreorderReasoning (⊆-preorder A) - open PreOrder public + open Base public hiding (step-≈; step-≈˘; step-∼) + renaming (∼-go to ⊆-go) - infixr 2 step-∼ step-⊆ - infix 1 step-∈ + open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → begin x) public + open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ ⊆-go public - step-⊆ = PreOrder.step-∼ + module _ {k : Related.ForwardKind} where + ∼-go : Trans _∼[ ⌊ k ⌋→ ]_ _IsRelatedTo_ _IsRelatedTo_ + ∼-go eq = ⊆-go (⇒→ eq) - step-∈ : ∀ x {xs ys : List A} → - xs IsRelatedTo ys → x ∈ xs → x ∈ ys - step-∈ x xs⊆ys x∈xs = (begin xs⊆ys) x∈xs + open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public - step-∼ : ∀ {k} xs {ys zs : List A} → - ys IsRelatedTo zs → xs ∼[ ⌊ k ⌋→ ] ys → xs IsRelatedTo zs - step-∼ xs ys⊆zs xs≈ys = step-⊆ xs ys⊆zs (⇒→ xs≈ys) - - syntax step-∈ x xs⊆ys x∈xs = x ∈⟨ x∈xs ⟩ xs⊆ys - syntax step-∼ xs ys⊆zs xs≈ys = xs ∼⟨ xs≈ys ⟩ ys⊆zs - syntax step-⊆ xs ys⊆zs xs⊆ys = xs ⊆⟨ xs⊆ys ⟩ ys⊆zs ------------------------------------------------------------------------ -- Congruence lemmas diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index aaea103279..28cf25d9ed 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -16,6 +16,7 @@ open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Transitive) open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) import Relation.Binary.Reasoning.Setoid as EqReasoning +open import Relation.Binary.Reasoning.Syntax ------------------------------------------------------------------------ -- An inductive definition of permutation @@ -73,16 +74,15 @@ data _↭_ : Rel (List A) a where module PermutationReasoning where - private - module Base = EqReasoning ↭-setoid + private module Base = EqReasoning ↭-setoid - open EqReasoning ↭-setoid public - hiding (step-≈; step-≈˘) + open Base public hiding (step-≈; step-≈˘) - infixr 2 step-↭ step-↭˘ step-swap step-prep + open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go ↭-sym public - step-↭ = Base.step-≈ - step-↭˘ = Base.step-≈˘ + -- Some extra combinators that allow us to skip certain elements + + infixr 2 step-swap step-prep -- Skip reasoning on the first element step-prep : ∀ x xs {ys zs : List A} → (x ∷ ys) IsRelatedTo zs → @@ -94,7 +94,5 @@ module PermutationReasoning where xs ↭ ys → (x ∷ y ∷ xs) IsRelatedTo zs step-swap x y xs rel xs↭ys = relTo (trans (swap x y xs↭ys) (begin rel)) - syntax step-↭ x y↭z x↭y = x ↭⟨ x↭y ⟩ y↭z - syntax step-↭˘ x y↭z y↭x = x ↭˘⟨ y↭x ⟩ y↭z syntax step-prep x xs y↭z x↭y = x ∷ xs <⟨ x↭y ⟩ y↭z syntax step-swap x y xs y↭z x↭y = x ∷ y ∷ xs <<⟨ x↭y ⟩ y↭z diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index fe3f095ee5..0d375226db 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -6,11 +6,13 @@ {-# OPTIONS --cubical-compatible --safe #-} +open import Function.Base using (_∘′_) open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive) +open import Relation.Binary.Reasoning.Syntax module Data.List.Relation.Binary.Permutation.Setoid {a ℓ} (S : Setoid a ℓ) where @@ -86,24 +88,16 @@ steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs module PermutationReasoning where - private - module Base = SetoidReasoning ↭-setoid + private module Base = SetoidReasoning ↭-setoid - open SetoidReasoning ↭-setoid public - hiding (step-≈; step-≈˘) + open Base public hiding (step-≈; step-≈˘) - infixr 2 step-↭ step-↭˘ step-≋ step-≋˘ step-swap step-prep + open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go ↭-sym public + open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ (Base.∼-go ∘′ refl) ≋-sym public - step-↭ = Base.step-≈ - step-↭˘ = Base.step-≈˘ + -- Some extra combinators that allow us to skip certain elements - -- Step with pointwise list equality - step-≋ : ∀ x {y z} → y IsRelatedTo z → x ≋ y → x IsRelatedTo z - step-≋ x (relTo y↔z) x≋y = relTo (trans (refl x≋y) y↔z) - - -- Step with flipped pointwise list equality - step-≋˘ : ∀ x {y z} → y IsRelatedTo z → y ≋ x → x IsRelatedTo z - step-≋˘ x y↭z y≋x = x ≋⟨ ≋-sym y≋x ⟩ y↭z + infixr 2 step-swap step-prep -- Skip reasoning on the first element step-prep : ∀ x xs {ys zs : List A} → (x ∷ ys) IsRelatedTo zs → @@ -115,9 +109,5 @@ module PermutationReasoning where xs ↭ ys → (x ∷ y ∷ xs) IsRelatedTo zs step-swap x y xs rel xs↭ys = relTo (trans (swap Eq.refl Eq.refl xs↭ys) (begin rel)) - syntax step-↭ x y↭z x↭y = x ↭⟨ x↭y ⟩ y↭z - syntax step-↭˘ x y↭z y↭x = x ↭˘⟨ y↭x ⟩ y↭z - syntax step-≋ x y↭z x≋y = x ≋⟨ x≋y ⟩ y↭z - syntax step-≋˘ x y↭z y≋x = x ≋˘⟨ y≋x ⟩ y↭z syntax step-prep x xs y↭z x↭y = x ∷ xs <⟨ x↭y ⟩ y↭z syntax step-swap x y xs y↭z x↭y = x ∷ y ∷ xs <<⟨ x↭y ⟩ y↭z diff --git a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda index abf5bdc86f..0d11079ad0 100644 --- a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda @@ -32,6 +32,7 @@ open import Relation.Binary.Definitions open import Relation.Binary.Bundles using (Setoid; Preorder) open import Relation.Binary.Structures using (IsPreorder) import Relation.Binary.Reasoning.Preorder as PreorderReasoning +open import Relation.Binary.Reasoning.Syntax open Setoid using (Carrier) @@ -112,31 +113,15 @@ module _ (S : Setoid a ℓ) where ------------------------------------------------------------------------ module ⊆-Reasoning (S : Setoid a ℓ) where + open Membership S using (_∈_) - open Setoid S renaming (Carrier to A) - open Subset S - open Membership S - - private - module Base = PreorderReasoning (⊆-preorder S) - - open Base public - hiding (step-∼; step-≈; step-≈˘) - - infixr 2 step-⊆ step-≋ step-≋˘ - infix 1 step-∈ - - step-∈ : ∀ x {xs ys} → xs IsRelatedTo ys → x ∈ xs → x ∈ ys - step-∈ x xs⊆ys x∈xs = (begin xs⊆ys) x∈xs + private module Base = PreorderReasoning (⊆-preorder S) - step-⊆ = Base.step-∼ - step-≋ = Base.step-≈ - step-≋˘ = Base.step-≈˘ + open Base public hiding (step-∼; step-≈; step-≈˘) - syntax step-∈ x xs⊆ys x∈xs = x ∈⟨ x∈xs ⟩ xs⊆ys - syntax step-⊆ xs ys⊆zs xs⊆ys = xs ⊆⟨ xs⊆ys ⟩ ys⊆zs - syntax step-≋ xs ys⊆zs xs≋ys = xs ≋⟨ xs≋ys ⟩ ys⊆zs - syntax step-≋˘ xs ys⊆zs xs≋ys = xs ≋˘⟨ xs≋ys ⟩ ys⊆zs + open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → Base.begin x) public + open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public + open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ Base.≈-go public ------------------------------------------------------------------------ -- Relationship with other binary relations diff --git a/src/Data/Nat/Binary/Properties.agda b/src/Data/Nat/Binary/Properties.agda index 8c08974636..9e8976842e 100644 --- a/src/Data/Nat/Binary/Properties.agda +++ b/src/Data/Nat/Binary/Properties.agda @@ -27,13 +27,7 @@ open import Function.Base using (_∘_; _$_; id) open import Function.Definitions open import Function.Consequences.Propositional open import Level using (0ℓ) -open import Relation.Binary.Core using (_⇒_; _Preserves_⟶_; _Preserves₂_⟶_⟶_) -open import Relation.Binary.Bundles - using (Setoid; DecSetoid; StrictPartialOrder; StrictTotalOrder; Preorder; Poset; TotalOrder; DecTotalOrder) -open import Relation.Binary.Definitions - using (Decidable; Irreflexive; Transitive; Reflexive; Antisymmetric; Total; Trichotomous; tri≈; tri<; tri>) -open import Relation.Binary.Structures - using (IsDecEquivalence; IsStrictPartialOrder; IsStrictTotalOrder; IsPreorder; IsPartialOrder; IsTotalOrder; IsDecTotalOrder) +open import Relation.Binary open import Relation.Binary.Consequences open import Relation.Binary.Morphism import Relation.Binary.Morphism.OrderMonomorphism as OrderMonomorphism @@ -371,6 +365,9 @@ toℕ-isMonomorphism-< = record <-trans (odd; tri≈; _Respectsʳ_; _Respectsˡ_; _Respects₂_) +open import Relation.Binary open import Relation.Binary.PropositionalEquality open import Relation.Binary.Morphism.Structures import Relation.Binary.Morphism.OrderMonomorphism as OrderMonomorphisms open import Relation.Nullary.Decidable.Core as Dec using (yes; no; recompute; map′; _×-dec_) open import Relation.Nullary.Negation.Core using (¬_; contradiction) +open import Relation.Binary.Reasoning.Syntax open import Algebra.Definitions {A = ℚ} _≡_ open import Algebra.Structures {A = ℚ} _≡_ @@ -136,11 +131,14 @@ mkℚ+-pos (suc n) (suc d) = _ -- Numerator and denominator equality ------------------------------------------------------------------------ +drop-*≡* : p ≃ q → ↥ p ℤ.* ↧ q ≡ ↥ q ℤ.* ↧ p +drop-*≡* (*≡* eq) = eq + ≡⇒≃ : _≡_ ⇒ _≃_ -≡⇒≃ refl = refl +≡⇒≃ refl = *≡* refl ≃⇒≡ : _≃_ ⇒ _≡_ -≃⇒≡ {x = mkℚ n₁ d₁ c₁} {y = mkℚ n₂ d₂ c₂} eq = helper +≃⇒≡ {x = mkℚ n₁ d₁ c₁} {y = mkℚ n₂ d₂ c₂} (*≡* eq) = helper where open ≡-Reasoning @@ -165,6 +163,9 @@ mkℚ+-pos (suc n) (suc d) = _ ... | refl with ℤ.*-cancelʳ-≡ n₁ n₂ (+ suc d₁) eq ... | refl = refl +≃-sym : Symmetric _≃_ +≃-sym = ≡⇒≃ ∘′ sym ∘′ ≃⇒≡ + ------------------------------------------------------------------------ -- Properties of ↥ ------------------------------------------------------------------------ @@ -176,6 +177,9 @@ mkℚ+-pos (suc n) (suc d) = _ p≡0⇒↥p≡0 : ∀ p → p ≡ 0ℚ → ↥ p ≡ 0ℤ p≡0⇒↥p≡0 p refl = refl +↥p≡↥q≡0⇒p≡q : ∀ p q → ↥ p ≡ 0ℤ → ↥ q ≡ 0ℤ → p ≡ q +↥p≡↥q≡0⇒p≡q p q ↥p≡0 ↥q≡0 = trans (↥p≡0⇒p≡0 p ↥p≡0) (sym (↥p≡0⇒p≡0 q ↥q≡0)) + ------------------------------------------------------------------------ -- Basic properties of sign predicates ------------------------------------------------------------------------ @@ -418,7 +422,7 @@ private ↧ᵘ-toℚᵘ p@record{} = refl toℚᵘ-injective : Injective _≡_ _≃ᵘ_ toℚᵘ -toℚᵘ-injective {x@record{}} {y@record{}} (*≡* eq) = ≃⇒≡ eq +toℚᵘ-injective {x@record{}} {y@record{}} (*≡* eq) = ≃⇒≡ (*≡* eq) fromℚᵘ-injective : Injective _≃ᵘ_ _≡_ fromℚᵘ fromℚᵘ-injective {p@record{}} {q@record{}} = /-injective-≃ p q @@ -497,7 +501,7 @@ private ≤-trans = ≤-Monomorphism.trans ℚᵘ.≤-trans ≤-antisym : Antisymmetric _≡_ _≤_ -≤-antisym (*≤* le₁) (*≤* le₂) = ≃⇒≡ (ℤ.≤-antisym le₁ le₂) +≤-antisym (*≤* le₁) (*≤* le₂) = ≃⇒≡ (*≡* (ℤ.≤-antisym le₁ le₂)) ≤-total : Total _≤_ ≤-total p q = [ inj₁ ∘ *≤* , inj₂ ∘ *≤* ]′ (ℤ.≤-total (↥ p ℤ.* ↧ q) (↥ q ℤ.* ↧ p)) @@ -606,7 +610,7 @@ toℚᵘ-isOrderMonomorphism-< = record ≰⇒> {p} {q} p≰q = *<* (ℤ.≰⇒> (p≰q ∘ *≤*)) <⇒≢ : _<_ ⇒ _≢_ -<⇒≢ {p} {q} (*<* p?_ = flip _ ≮ ≢ > = tri> (≮ ∘ drop-*<*) (≢ ∘ ≡⇒≃) (*<* >) +... | tri< < ≢ ≯ = tri< (*<* <) (≢ ∘ drop-*≡* ∘ ≡⇒≃) (≯ ∘ drop-*<*) +... | tri≈ ≮ ≡ ≯ = tri≈ (≮ ∘ drop-*<*) (≃⇒≡ (*≡* ≡)) (≯ ∘ drop-*<*) +... | tri> ≮ ≢ > = tri> (≮ ∘ drop-*<*) (≢ ∘ drop-*≡* ∘ ≡⇒≃) (*<* >) <-irrelevant : Irrelevant _<_ <-irrelevant (*<* p?_ = flip _?_ = flip _ Date: Tue, 17 Oct 2023 00:18:02 +0100 Subject: [PATCH 29/57] Fixes typos identified in #2154 (#2158) --- CHANGELOG.md | 52 +++++++++++++++++++-------------------- src/Function/Bundles.agda | 6 ++--- 2 files changed, 29 insertions(+), 29 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index efce415622..de399fe631 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -539,7 +539,7 @@ Non-backwards compatible changes 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. -* Similarly, in order to prevent reduction, the equality `_≃_` in `Data.Rational.Base` +* Similarly, in order to prevent reduction, the equality `_≃_` in `Data.Rational.Base` has been made into a data type with the single constructor `*≡*`. The destructor `drop-*≡*` has been added to `Data.Rational.Properties`. @@ -931,40 +931,40 @@ Non-backwards compatible changes as would be expected. Instead it omitted several fields like irreflexivity as they were derivable from the proof of trichotomy. However, this led to problems further up the hierarchy where bundles such as `StrictTotalOrder` which contained multiple distinct proofs of `IsStrictPartialOrder`. - -* To remedy this the definition of `IsStrictTotalOrder` has been changed to so that it builds upon - `IsStrictPartialOrder` as would be expected. + +* To remedy this the definition of `IsStrictTotalOrder` has been changed to so that it builds upon + `IsStrictPartialOrder` as would be expected. * To aid migration, the old record definition has been moved to `Relation.Binary.Structures.Biased` - which contains the `isStrictTotalOrderᶜ` smart constructor (which is re-exported by `Relation.Binary`) . + which contains the `isStrictTotalOrderᶜ` smart constructor (which is re-exported by `Relation.Binary`) . Therefore the old code: ```agda <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ <-isStrictTotalOrder = record - { isEquivalence = isEquivalence - ; trans = <-trans - ; compare = <-cmp - } + { isEquivalence = isEquivalence + ; trans = <-trans + ; compare = <-cmp + } ``` can be migrated either by updating to the new record fields if you have a proof of `IsStrictPartialOrder` available: ```agda <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ <-isStrictTotalOrder = record - { isStrictPartialOrder = <-isStrictPartialOrder - ; compare = <-cmp - } + { isStrictPartialOrder = <-isStrictPartialOrder + ; compare = <-cmp + } ``` or simply applying the smart constructor to the record definition as follows: ```agda <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ <-isStrictTotalOrder = isStrictTotalOrderᶜ record - { isEquivalence = isEquivalence - ; trans = <-trans - ; compare = <-cmp - } + { isEquivalence = isEquivalence + ; trans = <-trans + ; compare = <-cmp + } ``` - + ### Changes to triple reasoning interface * The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof that the strict @@ -1231,11 +1231,11 @@ Major improvements ### More modular design of equational reasoning. -* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports +* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports a range of modules containing pre-existing reasoning combinator syntax. -* This makes it possible to add new or rename existing reasoning combinators to a - pre-existing `Reasoning` module in just a couple of lines +* This makes it possible to add new or rename existing reasoning combinators to a + pre-existing `Reasoning` module in just a couple of lines (e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`) Deprecated modules @@ -1813,7 +1813,7 @@ Deprecated names ```agda _↔⟨⟩_ ↦ _≡⟨⟩_ ``` - + * In `Foreign.Haskell.Either` and `Foreign.Haskell.Pair`: ``` toForeign ↦ Foreign.Haskell.Coerce.coerce @@ -2711,7 +2711,7 @@ Additions to existing modules toℕ-inverseˡ : Inverseˡ _≡_ _≡_ toℕ fromℕ toℕ-inverseʳ : Inverseʳ _≡_ _≡_ toℕ fromℕ toℕ-inverseᵇ : Inverseᵇ _≡_ _≡_ toℕ fromℕ - + <-asym : Asymmetric _<_ ``` @@ -3168,7 +3168,7 @@ Additions to existing modules ∀ {m n} → _Respectsʳ_ (_<_ {m} {n}) _≋_ <-wellFounded : Transitive _≈_ → _≺_ Respectsʳ _≈_ → WellFounded _≺_ → ∀ {n} → WellFounded (_<_ {n}) -``` + ``` * Added new functions in `Data.Vec.Relation.Unary.Any`: ``` @@ -3195,9 +3195,9 @@ Additions to existing modules * Added new application operator synonym to `Function.Bundles`: ```agda _⟨$⟩_ : Func From To → Carrier From → Carrier To - _⟨$⟩_ = Func.to + _⟨$⟩_ = Func.to ``` - + * Added new proofs in `Function.Construct.Symmetry`: ``` bijective : Bijective ≈₁ ≈₂ f → Symmetric ≈₂ → Transitive ≈₂ → Congruent ≈₁ ≈₂ f → Bijective ≈₂ ≈₁ f⁻¹ @@ -3899,7 +3899,7 @@ This is a full list of proofs that have changed form to use irrelevant instance blockerAll : List Blocker → Blocker blockTC : Blocker → TC A ``` - + * Added new file `Relation.Binary.Reasoning.Base.Apartness` This is how to use it: diff --git a/src/Function/Bundles.agda b/src/Function/Bundles.agda index a51dc4ea84..0aef27fc57 100644 --- a/src/Function/Bundles.agda +++ b/src/Function/Bundles.agda @@ -182,7 +182,7 @@ module _ (From : Setoid a ℓ₁) (To : Setoid b ℓ₂) where open Func toFunction public using (module Eq₁; module Eq₂) - renaming (isCongruent to to-isCongrunet) + renaming (isCongruent to to-isCongruent) fromFunction : Func To From fromFunction = record @@ -192,7 +192,7 @@ module _ (From : Setoid a ℓ₁) (To : Setoid b ℓ₂) where open Func fromFunction public using () - renaming (isCongruent to from-isCongrunet) + renaming (isCongruent to from-isCongruent) record LeftInverse : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) where @@ -481,7 +481,7 @@ module _ {A : Set a} {B : Set b} where module _ {From : Setoid a ℓ₁} {To : Setoid b ℓ₂} where open Setoid - + infixl 5 _⟨$⟩_ _⟨$⟩_ : Func From To → Carrier From → Carrier To _⟨$⟩_ = Func.to From b0c95bca1716dbc243bfb3adb8fa7a71c777b66a Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Tue, 17 Oct 2023 00:18:55 +0100 Subject: [PATCH 30/57] tackles #2124 as regards `case_return_of_` (#2157) --- CHANGELOG.md | 5 +++++ README/Case.agda | 4 ++-- src/Function/Base.agda | 24 +++++++++++++++++++----- 3 files changed, 26 insertions(+), 7 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index de399fe631..233ea01257 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1776,6 +1776,11 @@ Deprecated names map-++-commute ↦ map-++ ``` +* In `Function.Base`: + ``` + case_return_of_ ↦ case_returning_of_ + ``` + * In `Function.Construct.Composition`: ``` _∘-⟶_ ↦ _⟶-∘_ diff --git a/README/Case.agda b/README/Case.agda index 30ce0221aa..62fec18c0f 100644 --- a/README/Case.agda +++ b/README/Case.agda @@ -12,7 +12,7 @@ module README.Case where open import Data.Fin hiding (pred) open import Data.Maybe hiding (from-just) open import Data.Nat hiding (pred) -open import Function.Base using (case_of_; case_return_of_) +open import Function.Base using (case_of_; case_returning_of_) open import Relation.Nullary ------------------------------------------------------------------------ @@ -35,7 +35,7 @@ pred n = case n of λ -- where-introduced and indentation-identified block of list of clauses from-just : ∀ {a} {A : Set a} (x : Maybe A) → From-just x -from-just x = case x return From-just of λ where +from-just x = case x returning From-just of λ where (just x) → x nothing → _ diff --git a/src/Function/Base.agda b/src/Function/Base.agda index 9bdb996dc8..122e276469 100644 --- a/src/Function/Base.agda +++ b/src/Function/Base.agda @@ -110,10 +110,10 @@ f $- = f _ -- Case expressions (to be used with pattern-matching lambdas, see -- README.Case). -case_return_of_ : ∀ {A : Set a} (x : A) (B : A → Set b) → +case_returning_of_ : ∀ {A : Set a} (x : A) (B : A → Set b) → ((x : A) → B x) → B x -case x return B of f = f x -{-# INLINE case_return_of_ #-} +case x returning B of f = f x +{-# INLINE case_returning_of_ #-} ------------------------------------------------------------------------ -- Non-dependent versions of dependent operations @@ -156,7 +156,7 @@ _|>′_ = _|>_ -- README.Case). case_of_ : A → (A → B) → B -case x of f = case x return _ of f +case x of f = case x returning _ of f {-# INLINE case_of_ #-} ------------------------------------------------------------------------ @@ -247,10 +247,24 @@ _*_ on f = f -⟨ _*_ ⟩- f ------------------------------------------------------------------------ --- Deprecated +-- DEPRECATED NAMES +------------------------------------------------------------------------ +-- Please use the new names as continuing support for the old names is +-- not guaranteed. + +-- Version 1.4 _-[_]-_ = _-⟪_⟫-_ {-# WARNING_ON_USAGE _-[_]-_ "Warning: Function._-[_]-_ was deprecated in v1.4. Please use _-⟪_⟫-_ instead." #-} + +-- Version 2.0 + +case_return_of_ = case_returning_of_ +{-# WARNING_ON_USAGE case_return_of_ +"case_return_of_ was deprecated in v2.0. +Please use case_returning_of_ instead." +#-} + From ef66f77ad3f5c2373abd5d9659002c7a4ff5812e Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Tue, 17 Oct 2023 08:19:32 +0900 Subject: [PATCH 31/57] Rename preorder ~ relation reasoning combinators (#2156) --- src/Codata/Musical/Colist.agda | 6 ++-- src/Data/Integer/Divisibility/Signed.agda | 6 ++-- .../Relation/Binary/BagAndSetEquality.agda | 4 +-- .../Binary/Subset/Setoid/Properties.agda | 8 +++-- src/Data/Nat/Divisibility.agda | 6 ++-- src/Data/Rational/Properties.agda | 4 +-- .../Rational/Unnormalised/Properties.agda | 2 +- .../ReflexiveTransitive/Properties.agda | 8 +++-- .../Binary/Reasoning/Base/Double.agda | 33 ++++++++++++++----- src/Relation/Binary/Reasoning/Syntax.agda | 14 +++++--- 10 files changed, 61 insertions(+), 30 deletions(-) diff --git a/src/Codata/Musical/Colist.agda b/src/Codata/Musical/Colist.agda index aeff3076dc..3699ef64be 100644 --- a/src/Codata/Musical/Colist.agda +++ b/src/Codata/Musical/Colist.agda @@ -190,10 +190,12 @@ module ⊑-Reasoning {a} {A : Set a} where module ⊆-Reasoning {A : Set a} where private module Base = PreR (⊆-Preorder A) - open Base public hiding (step-∼) + open Base public + hiding (step-≲; step-∼) + renaming (≲-go to ⊆-go) open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → begin x) - open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public + open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ ⊆-go public -- take returns a prefix. diff --git a/src/Data/Integer/Divisibility/Signed.agda b/src/Data/Integer/Divisibility/Signed.agda index 9a4b938950..5149871534 100644 --- a/src/Data/Integer/Divisibility/Signed.agda +++ b/src/Data/Integer/Divisibility/Signed.agda @@ -121,9 +121,11 @@ open _∣_ using (quotient) public module ∣-Reasoning where private module Base = PreorderReasoning ∣-preorder - open Base public hiding (step-∼; step-≈; step-≈˘) + open Base public + hiding (step-≲; step-∼; step-≈; step-≈˘) + renaming (≲-go to ∣-go) - open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public + open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ ∣-go public ------------------------------------------------------------------------ -- Other properties of _∣_ diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 0acef52dcb..f842ae8087 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -96,8 +96,8 @@ module ⊆-Reasoning {A : Set a} where private module Base = PreorderReasoning (⊆-preorder A) open Base public - hiding (step-≈; step-≈˘; step-∼) - renaming (∼-go to ⊆-go) + hiding (step-≈; step-≈˘; step-∼; step-≲) + renaming (≲-go to ⊆-go) open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → begin x) public open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ ⊆-go public diff --git a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda index 0d11079ad0..bb20fd417f 100644 --- a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda @@ -117,11 +117,13 @@ module ⊆-Reasoning (S : Setoid a ℓ) where private module Base = PreorderReasoning (⊆-preorder S) - open Base public hiding (step-∼; step-≈; step-≈˘) + open Base public + hiding (step-≲; step-≈; step-≈˘; step-∼) + renaming (≲-go to ⊆-go; ≈-go to ≋-go) open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → Base.begin x) public - open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public - open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ Base.≈-go public + open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ ⊆-go public + open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ ≋-go public ------------------------------------------------------------------------ -- Relationship with other binary relations diff --git a/src/Data/Nat/Divisibility.agda b/src/Data/Nat/Divisibility.agda index 5a0491e498..63037a49b2 100644 --- a/src/Data/Nat/Divisibility.agda +++ b/src/Data/Nat/Divisibility.agda @@ -120,9 +120,11 @@ suc n ∣? m = Dec.map (m%n≡0⇔n∣m m (suc n)) (m % suc n ≟ 0) module ∣-Reasoning where private module Base = PreorderReasoning ∣-preorder - open Base public hiding (step-≈; step-≈˘; step-∼) + open Base public + hiding (step-≈; step-≈˘; step-∼; step-≲) + renaming (≲-go to ∣-go) - open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public + open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ ∣-go public ------------------------------------------------------------------------ -- Simple properties of _∣_ diff --git a/src/Data/Rational/Properties.agda b/src/Data/Rational/Properties.agda index b43b109943..915eb6e058 100644 --- a/src/Data/Rational/Properties.agda +++ b/src/Data/Rational/Properties.agda @@ -164,7 +164,7 @@ drop-*≡* (*≡* eq) = eq ... | refl = refl ≃-sym : Symmetric _≃_ -≃-sym = ≡⇒≃ ∘′ sym ∘′ ≃⇒≡ +≃-sym = ≡⇒≃ ∘′ sym ∘′ ≃⇒≡ ------------------------------------------------------------------------ -- Properties of ↥ @@ -753,7 +753,7 @@ module ≤-Reasoning where ≃-go : Trans _≃_ _IsRelatedTo_ _IsRelatedTo_ ≃-go = Triple.≈-go ∘′ ≃⇒≡ - + open ≃-syntax _IsRelatedTo_ _IsRelatedTo_ ≃-go (λ {x y} → ≃-sym {x} {y}) public ------------------------------------------------------------------------ diff --git a/src/Data/Rational/Unnormalised/Properties.agda b/src/Data/Rational/Unnormalised/Properties.agda index dffd2529b0..6348a6ef46 100644 --- a/src/Data/Rational/Unnormalised/Properties.agda +++ b/src/Data/Rational/Unnormalised/Properties.agda @@ -164,7 +164,7 @@ p≃0⇒↥p≡0 p (*≡* eq) = begin ↥ p ℤ.* 1ℤ ≡⟨ eq ⟩ 0ℤ ∎ where open ≡-Reasoning - + ↥p≡↥q≡0⇒p≃q : ∀ p q → ↥ p ≡ 0ℤ → ↥ q ≡ 0ℤ → p ≃ q ↥p≡↥q≡0⇒p≃q p q ↥p≡0 ↥q≡0 = ≃-trans (↥p≡0⇒p≃0 p ↥p≡0) (≃-sym (↥p≡0⇒p≃0 _ ↥q≡0)) diff --git a/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda b/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda index c8ebd64b7c..58b2ee198a 100644 --- a/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda +++ b/src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda @@ -115,7 +115,9 @@ module _ {i t} {I : Set i} (T : Rel I t) where module StarReasoning {i t} {I : Set i} (T : Rel I t) where private module Base = PreorderReasoning (preorder T) - open Base public hiding (step-≈; step-∼) + open Base public + hiding (step-≈; step-∼; step-≲) + renaming (≲-go to ⟶-go) - open ⟶-syntax _IsRelatedTo_ _IsRelatedTo_ (Base.∼-go ∘ (_◅ ε)) public - open ⟶*-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public + open ⟶-syntax _IsRelatedTo_ _IsRelatedTo_ (⟶-go ∘ (_◅ ε)) public + open ⟶*-syntax _IsRelatedTo_ _IsRelatedTo_ ⟶-go public diff --git a/src/Relation/Binary/Reasoning/Base/Double.agda b/src/Relation/Binary/Reasoning/Base/Double.agda index 91203437f6..6fa901d573 100644 --- a/src/Relation/Binary/Reasoning/Base/Double.agda +++ b/src/Relation/Binary/Reasoning/Base/Double.agda @@ -21,7 +21,7 @@ open import Relation.Binary.Reasoning.Syntax module Relation.Binary.Reasoning.Base.Double {a ℓ₁ ℓ₂} {A : Set a} - {_≈_ : Rel A ℓ₁} {_∼_ : Rel A ℓ₂} (isPreorder : IsPreorder _≈_ _∼_) + {_≈_ : Rel A ℓ₁} {_≲_ : Rel A ℓ₂} (isPreorder : IsPreorder _≈_ _≲_) where open IsPreorder isPreorder @@ -32,24 +32,24 @@ open IsPreorder isPreorder infix 4 _IsRelatedTo_ data _IsRelatedTo_ (x y : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where - nonstrict : (x∼y : x ∼ y) → x IsRelatedTo y + nonstrict : (x≲y : x ≲ y) → x IsRelatedTo y equals : (x≈y : x ≈ y) → x IsRelatedTo y -start : _IsRelatedTo_ ⇒ _∼_ +start : _IsRelatedTo_ ⇒ _≲_ start (equals x≈y) = reflexive x≈y -start (nonstrict x∼y) = x∼y +start (nonstrict x≲y) = x≲y ≡-go : Trans _≡_ _IsRelatedTo_ _IsRelatedTo_ ≡-go x≡y (equals y≈z) = equals (case x≡y of λ where P.refl → y≈z) ≡-go x≡y (nonstrict y≤z) = nonstrict (case x≡y of λ where P.refl → y≤z) -∼-go : Trans _∼_ _IsRelatedTo_ _IsRelatedTo_ -∼-go x∼y (equals y≈z) = nonstrict (∼-respʳ-≈ y≈z x∼y) -∼-go x∼y (nonstrict y∼z) = nonstrict (trans x∼y y∼z) +≲-go : Trans _≲_ _IsRelatedTo_ _IsRelatedTo_ +≲-go x≲y (equals y≈z) = nonstrict (∼-respʳ-≈ y≈z x≲y) +≲-go x≲y (nonstrict y≲z) = nonstrict (trans x≲y y≲z) ≈-go : Trans _≈_ _IsRelatedTo_ _IsRelatedTo_ ≈-go x≈y (equals y≈z) = equals (Eq.trans x≈y y≈z) -≈-go x≈y (nonstrict y∼z) = nonstrict (∼-respˡ-≈ (Eq.sym x≈y) y∼z) +≈-go x≈y (nonstrict y≲z) = nonstrict (∼-respˡ-≈ (Eq.sym x≈y) y≲z) stop : Reflexive _IsRelatedTo_ stop = equals Eq.refl @@ -81,6 +81,21 @@ equalitySubRelation = record open begin-syntax _IsRelatedTo_ start public open begin-equality-syntax _IsRelatedTo_ equalitySubRelation public open ≡-syntax _IsRelatedTo_ ≡-go public -open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public open ≈-syntax _IsRelatedTo_ _IsRelatedTo_ ≈-go Eq.sym public +open ≲-syntax _IsRelatedTo_ _IsRelatedTo_ ≲-go public open end-syntax _IsRelatedTo_ stop public + + +------------------------------------------------------------------------ +-- DEPRECATED NAMES +------------------------------------------------------------------------ +-- Please use the new names as continuing support for the old names is +-- not guaranteed. + +-- Version 2.0 + +open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ≲-go public +{-# WARNING_ON_USAGE step-∼ +"Warning: step-∼ and _∼⟨_⟩_ syntax was deprecated in v2.0. +Please use step-≲ and _≲⟨_⟩_ instead. " +#-} diff --git a/src/Relation/Binary/Reasoning/Syntax.agda b/src/Relation/Binary/Reasoning/Syntax.agda index 8bd7899770..4837fe33f4 100644 --- a/src/Relation/Binary/Reasoning/Syntax.agda +++ b/src/Relation/Binary/Reasoning/Syntax.agda @@ -1,7 +1,7 @@ ------------------------------------------------------------------------ -- The Agda standard library -- --- Syntax for the building blocks of equational reasoning modules +-- Syntax for the building blocks of equational reasoning modules ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} @@ -26,7 +26,7 @@ open import Relation.Binary.PropositionalEquality.Core as P -- Codata/Guarded/Stream/Relation/Binary/Pointwise -- Codata/Sized/Stream/Bisimilarity -- Function/Reasoning - + module Relation.Binary.Reasoning.Syntax where private @@ -170,14 +170,20 @@ module _ forward : ∀ (x : A) {y z} → S y z → R x y → T x z forward x yRz x∼y = step {x} x∼y yRz - - -- Preorder syntax + -- Arbitrary relation syntax module ∼-syntax where infixr 2 step-∼ step-∼ = forward syntax step-∼ x yRz x∼y = x ∼⟨ x∼y ⟩ yRz + -- Preorder syntax + module ≲-syntax where + infixr 2 step-≲ + step-≲ = forward + syntax step-≲ x yRz x≲y = x ≲⟨ x≲y ⟩ yRz + + -- Partial order syntax module ≤-syntax where infixr 2 step-≤ From 33811e50e3ecac3f93e312d023e44995d535b94f Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Wed, 18 Oct 2023 07:20:25 +0900 Subject: [PATCH 32/57] =?UTF-8?q?Move=20=E2=89=A1-Reasoning=20from=20Core?= =?UTF-8?q?=20to=20Properties=20and=20implement=20using=20syntax=20(#2159)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- CHANGELOG.md | 6 ++++ src/Codata/Guarded/Stream/Properties.agda | 6 ++-- src/Codata/Musical/Colist/Infinite-merge.agda | 4 ++- src/Codata/Sized/Stream/Properties.agda | 4 ++- src/Data/Fin/Relation/Unary/Top.agda | 2 ++ src/Data/Fin/Substitution/Lemmas.agda | 4 ++- src/Data/List/Countdown.agda | 2 +- .../Example/UniqueBoundVariables.agda | 3 +- .../Ternary/Interleaving/Properties.agda | 4 ++- .../List/Relation/Unary/Any/Properties.agda | 4 ++- src/Data/Nat/Coprimality.agda | 2 +- src/Data/Nat/GCD.agda | 4 ++- src/Data/Nat/LCM.agda | 4 ++- src/Data/String/Unsafe.agda | 5 ++- src/Data/Vec/Bounded/Base.agda | 4 ++- src/Data/Vec/Recursive/Properties.agda | 2 ++ .../Vec/Relation/Binary/Equality/Cast.agda | 4 ++- src/Effect/Applicative/Indexed.agda | 4 ++- src/Function/Related/TypeIsomorphisms.agda | 4 ++- .../Binary/PropositionalEquality/Core.agda | 32 ------------------- .../PropositionalEquality/Properties.agda | 15 +++++++++ src/Relation/Binary/Reasoning/Syntax.agda | 1 - 22 files changed, 70 insertions(+), 50 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 233ea01257..0ba8b736c4 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1238,6 +1238,12 @@ Major improvements pre-existing `Reasoning` module in just a couple of lines (e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`) +* One pre-requisite for that is that `≡-Reasoning` has been moved from + `Relation.Binary.PropositionalEquality.Core` (which shouldn't be + imported anyway as it's a `Core` module) to + `Relation.Binary.PropositionalEquality.Properties`. + It is still exported by `Relation.Binary.PropositionalEquality`. + Deprecated modules ------------------ diff --git a/src/Codata/Guarded/Stream/Properties.agda b/src/Codata/Guarded/Stream/Properties.agda index bd41d1967e..3742aab856 100644 --- a/src/Codata/Guarded/Stream/Properties.agda +++ b/src/Codata/Guarded/Stream/Properties.agda @@ -21,6 +21,8 @@ open import Data.Vec.Base as Vec using (Vec; _∷_) open import Function.Base using (const; flip; id; _∘′_; _$′_; _⟨_⟩_; _∘₂′_) open import Level using (Level) open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; cong; cong₂) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) private variable @@ -218,7 +220,7 @@ lookup-transpose n (as ∷ ass) = begin lookup as n ∷ lookup (transpose ass) n ≡⟨ cong (lookup as n ∷_) (lookup-transpose n ass) ⟩ lookup as n ∷ List.map (flip lookup n) ass ≡⟨⟩ List.map (flip lookup n) (as ∷ ass) ∎ - where open P.≡-Reasoning + where open ≡-Reasoning lookup-transpose⁺ : ∀ n (ass : List⁺ (Stream A)) → lookup (transpose⁺ ass) n ≡ List⁺.map (flip lookup n) ass @@ -228,7 +230,7 @@ lookup-transpose⁺ n (as ∷ ass) = begin lookup as n ∷ lookup (transpose ass) n ≡⟨ cong (lookup as n ∷_) (lookup-transpose n ass) ⟩ lookup as n ∷ List.map (flip lookup n) ass ≡⟨⟩ List⁺.map (flip lookup n) (as ∷ ass) ∎ - where open P.≡-Reasoning + where open ≡-Reasoning lookup-tails : ∀ n (as : Stream A) → lookup (tails as) n ≈ ℕ.iterate tail as n lookup-tails zero as = B.refl diff --git a/src/Codata/Musical/Colist/Infinite-merge.agda b/src/Codata/Musical/Colist/Infinite-merge.agda index cbc142ff5a..584880a768 100644 --- a/src/Codata/Musical/Colist/Infinite-merge.agda +++ b/src/Codata/Musical/Colist/Infinite-merge.agda @@ -26,6 +26,8 @@ open import Level open import Relation.Unary using (Pred) import Induction.WellFounded as WF open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) import Relation.Binary.Construct.On as On private @@ -132,7 +134,7 @@ merge xss = ⟦ merge′ xss ⟧P Any-merge : ∀ xss → Any P (merge xss) ↔ Any (λ { (x , xs) → P x ⊎ Any P xs }) xss Any-merge {P = P} xss = mk↔ₛ′ (proj₁ ∘ to xss) from to∘from (proj₂ ∘ to xss) where - open P.≡-Reasoning + open ≡-Reasoning -- The from function. diff --git a/src/Codata/Sized/Stream/Properties.agda b/src/Codata/Sized/Stream/Properties.agda index d2cf4e3205..40cb3fc961 100644 --- a/src/Codata/Sized/Stream/Properties.agda +++ b/src/Codata/Sized/Stream/Properties.agda @@ -25,6 +25,8 @@ open import Data.Vec.Base as Vec using (_∷_) open import Function.Base using (id; _$_; _∘′_; const) open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; _≢_) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) private variable @@ -116,7 +118,7 @@ lookup-iterate-identity (suc n) f a = begin lookup (iterate f (f a)) n ≡⟨ lookup-iterate-identity n f (f a) ⟩ fold (f a) f n ≡⟨ fold-pull a f (const ∘′ f) (f a) P.refl (λ _ → P.refl) n ⟩ f (fold a f n) ≡⟨⟩ - fold a f (suc n) ∎ where open P.≡-Reasoning + fold a f (suc n) ∎ where open ≡-Reasoning ------------------------------------------------------------------------ -- DEPRECATED diff --git a/src/Data/Fin/Relation/Unary/Top.agda b/src/Data/Fin/Relation/Unary/Top.agda index b6a567a9a8..a884a09243 100644 --- a/src/Data/Fin/Relation/Unary/Top.agda +++ b/src/Data/Fin/Relation/Unary/Top.agda @@ -16,6 +16,8 @@ module Data.Fin.Relation.Unary.Top where open import Data.Nat.Base using (ℕ; zero; suc) open import Data.Fin.Base using (Fin; zero; suc; fromℕ; inject₁) open import Relation.Binary.PropositionalEquality.Core +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) private variable diff --git a/src/Data/Fin/Substitution/Lemmas.agda b/src/Data/Fin/Substitution/Lemmas.agda index ffa26e64fb..8b4d0b176a 100644 --- a/src/Data/Fin/Substitution/Lemmas.agda +++ b/src/Data/Fin/Substitution/Lemmas.agda @@ -16,9 +16,11 @@ import Data.Vec.Properties as VecProp open import Function.Base as Fun using (_∘_; _$_; flip) open import Relation.Binary.PropositionalEquality.Core as PropEq using (_≡_; refl; sym; cong; cong₂) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open import Relation.Binary.Construct.Closure.ReflexiveTransitive using (Star; ε; _◅_; _▻_) -open PropEq.≡-Reasoning +open ≡-Reasoning open import Level using (Level; _⊔_) open import Relation.Unary using (Pred) diff --git a/src/Data/List/Countdown.agda b/src/Data/List/Countdown.agda index 8d6761de27..dbfd75b332 100644 --- a/src/Data/List/Countdown.agda +++ b/src/Data/List/Countdown.agda @@ -31,8 +31,8 @@ open import Relation.Nullary open import Relation.Nullary.Decidable using (dec-true; dec-false) open import Relation.Binary.PropositionalEquality.Core as PropEq using (_≡_; _≢_; refl; cong) -open PropEq.≡-Reasoning import Relation.Binary.PropositionalEquality.Properties as PropEq +open PropEq.≡-Reasoning private open module D = DecSetoid D diff --git a/src/Data/List/Relation/Binary/Sublist/Propositional/Example/UniqueBoundVariables.agda b/src/Data/List/Relation/Binary/Sublist/Propositional/Example/UniqueBoundVariables.agda index 39a6fe65a1..163818b6e8 100644 --- a/src/Data/List/Relation/Binary/Sublist/Propositional/Example/UniqueBoundVariables.agda +++ b/src/Data/List/Relation/Binary/Sublist/Propositional/Example/UniqueBoundVariables.agda @@ -10,7 +10,8 @@ module Data.List.Relation.Binary.Sublist.Propositional.Example.UniqueBoundVariables (Base : Set) where -open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; sym; cong; subst; module ≡-Reasoning) +open import Relation.Binary.PropositionalEquality + using (_≡_; refl; sym; cong; subst; module ≡-Reasoning) open ≡-Reasoning open import Data.List.Base using (List; []; _∷_; [_]) diff --git a/src/Data/List/Relation/Ternary/Interleaving/Properties.agda b/src/Data/List/Relation/Ternary/Interleaving/Properties.agda index dff26bb6b2..5317d5a7c5 100644 --- a/src/Data/List/Relation/Ternary/Interleaving/Properties.agda +++ b/src/Data/List/Relation/Ternary/Interleaving/Properties.agda @@ -16,7 +16,9 @@ open import Data.List.Relation.Ternary.Interleaving hiding (map) open import Function.Base using (_$_) open import Relation.Binary.Core using (REL) open import Relation.Binary.PropositionalEquality.Core - using (_≡_; refl; sym; cong; module ≡-Reasoning) + using (_≡_; refl; sym; cong) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open ≡-Reasoning ------------------------------------------------------------------------ diff --git a/src/Data/List/Relation/Unary/Any/Properties.agda b/src/Data/List/Relation/Unary/Any/Properties.agda index aba1a777ee..b88bdfbe95 100644 --- a/src/Data/List/Relation/Unary/Any/Properties.agda +++ b/src/Data/List/Relation/Unary/Any/Properties.agda @@ -43,6 +43,8 @@ open import Relation.Binary.Core using (Rel; REL) open import Relation.Binary.Definitions as B open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open import Relation.Unary as U using (Pred; _⟨×⟩_; _⟨→⟩_) renaming (_⊆_ to _⋐_) open import Relation.Nullary using (¬_; _because_; does; ofʸ; ofⁿ; yes; no) @@ -219,7 +221,7 @@ Any-×⁻ pq with Prod.map₂ (Prod.map₂ find) (find pq) (Any P xs × Any Q ys) ↔ Any (λ x → Any (λ y → P x × Q y) ys) xs ×↔ {P = P} {Q = Q} {xs} {ys} = mk↔ₛ′ Any-×⁺ Any-×⁻ to∘from from∘to where - open P.≡-Reasoning + open ≡-Reasoning from∘to : ∀ pq → Any-×⁻ (Any-×⁺ pq) ≡ pq from∘to (p , q) = diff --git a/src/Data/Nat/Coprimality.agda b/src/Data/Nat/Coprimality.agda index f6f91b44e5..1df320ea96 100644 --- a/src/Data/Nat/Coprimality.agda +++ b/src/Data/Nat/Coprimality.agda @@ -20,7 +20,7 @@ open import Data.Product.Base as Prod open import Function.Base using (_∘_) open import Level using (0ℓ) open import Relation.Binary.PropositionalEquality.Core as P - using (_≡_; _≢_; refl; trans; cong; subst; module ≡-Reasoning) + using (_≡_; _≢_; refl; trans; cong; subst) open import Relation.Nullary as Nullary hiding (recompute) open import Relation.Nullary.Negation using (contradiction) open import Relation.Binary.Core using (Rel) diff --git a/src/Data/Nat/GCD.agda b/src/Data/Nat/GCD.agda index 555a902df6..b88947c707 100644 --- a/src/Data/Nat/GCD.agda +++ b/src/Data/Nat/GCD.agda @@ -24,6 +24,8 @@ open import Induction.Lexicographic using (_⊗_; [_⊗_]) open import Relation.Binary.Definitions using (tri<; tri>; tri≈; Symmetric) open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; _≢_; subst; cong) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open import Relation.Nullary.Decidable using (Dec) open import Relation.Nullary.Negation using (contradiction) import Relation.Nullary.Decidable as Dec @@ -190,7 +192,7 @@ c*gcd[m,n]≡gcd[cm,cn] c@(suc _) m n = begin c * gcd m n ≡⟨ cong (c *_) (P.sym (gcd[cm,cn]/c≡gcd[m,n] c m n)) ⟩ c * (gcd (c * m) (c * n) / c) ≡⟨ m*[n/m]≡n (gcd-greatest (m∣m*n m) (m∣m*n n)) ⟩ gcd (c * m) (c * n) ∎ - where open P.≡-Reasoning + where open ≡-Reasoning gcd[m,n]≤n : ∀ m n .{{_ : NonZero n}} → gcd m n ≤ n gcd[m,n]≤n m n = ∣⇒≤ (gcd[m,n]∣n m n) diff --git a/src/Data/Nat/LCM.agda b/src/Data/Nat/LCM.agda index 545870574f..5dff2eb530 100644 --- a/src/Data/Nat/LCM.agda +++ b/src/Data/Nat/LCM.agda @@ -18,7 +18,9 @@ open import Data.Nat.GCD open import Data.Product.Base using (_×_; _,_; uncurry′; ∃) open import Data.Sum.Base using (_⊎_; inj₁; inj₂) open import Relation.Binary.PropositionalEquality.Core as P - using (_≡_; refl; sym; trans; cong; cong₂; module ≡-Reasoning) + using (_≡_; refl; sym; trans; cong; cong₂) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open import Relation.Nullary.Decidable using (False; fromWitnessFalse) private diff --git a/src/Data/String/Unsafe.agda b/src/Data/String/Unsafe.agda index ff1ff28408..35b408f131 100644 --- a/src/Data/String/Unsafe.agda +++ b/src/Data/String/Unsafe.agda @@ -16,8 +16,11 @@ open import Data.Product.Base using (proj₂) open import Data.String.Base open import Function.Base using (_∘′_) -open import Relation.Binary.PropositionalEquality.Core; open ≡-Reasoning +open import Relation.Binary.PropositionalEquality.Core +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open import Relation.Binary.PropositionalEquality.TrustMe using (trustMe) +open ≡-Reasoning ------------------------------------------------------------------------ -- Properties of tail diff --git a/src/Data/Vec/Bounded/Base.agda b/src/Data/Vec/Bounded/Base.agda index 63649e940e..c926c9dfb1 100644 --- a/src/Data/Vec/Bounded/Base.agda +++ b/src/Data/Vec/Bounded/Base.agda @@ -22,6 +22,8 @@ open import Function.Base using (_∘_; id; _$_) open import Level using (Level) open import Relation.Nullary.Decidable.Core using (recompute) open import Relation.Binary.PropositionalEquality.Core as P using (_≡_; refl) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) private variable @@ -66,7 +68,7 @@ private m + (⌈ k /2⌉ + ⌊ k /2⌋) ≡⟨ P.cong (m +_) (ℕₚ.+-comm ⌈ k /2⌉ ⌊ k /2⌋) ⟩ m + (⌊ k /2⌋ + ⌈ k /2⌉) ≡⟨ P.cong (m +_) (ℕₚ.⌊n/2⌋+⌈n/2⌉≡n k) ⟩ m + k ≡⟨ eq ⟩ - n ∎ where open P.≡-Reasoning + n ∎ where open ≡-Reasoning padBoth : ∀ {n} → A → A → Vec≤ A n → Vec A n padBoth aₗ aᵣ as@(vs , m≤n) diff --git a/src/Data/Vec/Recursive/Properties.agda b/src/Data/Vec/Recursive/Properties.agda index b012e59a80..d036134a45 100644 --- a/src/Data/Vec/Recursive/Properties.agda +++ b/src/Data/Vec/Recursive/Properties.agda @@ -15,6 +15,8 @@ open import Data.Vec.Recursive open import Data.Vec.Base using (Vec; _∷_) open import Function.Bundles using (_↔_; mk↔ₛ′) open import Relation.Binary.PropositionalEquality.Core as P +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open ≡-Reasoning private diff --git a/src/Data/Vec/Relation/Binary/Equality/Cast.agda b/src/Data/Vec/Relation/Binary/Equality/Cast.agda index adca0a42f2..f65981f7a3 100644 --- a/src/Data/Vec/Relation/Binary/Equality/Cast.agda +++ b/src/Data/Vec/Relation/Binary/Equality/Cast.agda @@ -18,7 +18,9 @@ open import Data.Vec.Base open import Relation.Binary.Core using (REL; _⇒_) open import Relation.Binary.Definitions using (Sym; Trans) open import Relation.Binary.PropositionalEquality.Core - using (_≡_; refl; trans; sym; cong; module ≡-Reasoning) + using (_≡_; refl; trans; sym; cong) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) private variable diff --git a/src/Effect/Applicative/Indexed.agda b/src/Effect/Applicative/Indexed.agda index 30e9b632bb..f547e68b27 100644 --- a/src/Effect/Applicative/Indexed.agda +++ b/src/Effect/Applicative/Indexed.agda @@ -16,6 +16,8 @@ open import Data.Product.Base using (_×_; _,_) open import Function.Base open import Level open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) private variable @@ -111,4 +113,4 @@ record Morphism {I : Set i} {F₁ F₂ : IFun I f} op (A₁._⊛_ (A₁.pure f) x) ≡⟨ op-⊛ _ _ ⟩ A₂._⊛_ (op (A₁.pure f)) (op x) ≡⟨ P.cong₂ A₂._⊛_ (op-pure _) P.refl ⟩ A₂._⊛_ (A₂.pure f) (op x) ∎ - where open P.≡-Reasoning + where open ≡-Reasoning diff --git a/src/Function/Related/TypeIsomorphisms.agda b/src/Function/Related/TypeIsomorphisms.agda index 830bb6f8bb..9fb1fded28 100644 --- a/src/Function/Related/TypeIsomorphisms.agda +++ b/src/Function/Related/TypeIsomorphisms.agda @@ -29,6 +29,8 @@ open import Function.Related.Propositional import Function.Construct.Identity as Identity open import Relation.Binary hiding (_⇔_) open import Relation.Binary.PropositionalEquality.Core as P using (_≡_) +open import Relation.Binary.PropositionalEquality.Properties + using (module ≡-Reasoning) open import Relation.Nullary.Reflects using (invert) open import Relation.Nullary using (Dec; ¬_; _because_; ofⁿ) import Relation.Nullary.Indexed as I @@ -284,7 +286,7 @@ private from C↔D (to C↔D (f (from A↔B (to A↔B x)))) ≡⟨ strictlyInverseʳ C↔D _ ⟩ f (from A↔B (to A↔B x)) ≡⟨ P.cong f $ strictlyInverseʳ A↔B x ⟩ f x ∎) - where open Inverse; open P.≡-Reasoning + where open Inverse; open ≡-Reasoning →-cong : Extensionality a c → Extensionality b d → ∀ {k} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → diff --git a/src/Relation/Binary/PropositionalEquality/Core.agda b/src/Relation/Binary/PropositionalEquality/Core.agda index c0c4030b83..848258348b 100644 --- a/src/Relation/Binary/PropositionalEquality/Core.agda +++ b/src/Relation/Binary/PropositionalEquality/Core.agda @@ -91,35 +91,3 @@ resp₂ _∼_ = respʳ _∼_ , respˡ _∼_ ≢-sym : Symmetric {A = A} _≢_ ≢-sym x≢y = x≢y ∘ sym - ------------------------------------------------------------------------- --- Convenient syntax for equational reasoning - --- This is a special instance of `Relation.Binary.Reasoning.Setoid`. --- Rather than instantiating the latter with (setoid A), we reimplement --- equation chains from scratch since then goals are printed much more --- readably. - -module ≡-Reasoning {A : Set a} where - - infix 3 _∎ - infixr 2 _≡⟨⟩_ step-≡ step-≡˘ - infix 1 begin_ - - begin_ : ∀{x y : A} → x ≡ y → x ≡ y - begin_ x≡y = x≡y - - _≡⟨⟩_ : ∀ (x {y} : A) → x ≡ y → x ≡ y - _ ≡⟨⟩ x≡y = x≡y - - step-≡ : ∀ (x {y z} : A) → y ≡ z → x ≡ y → x ≡ z - step-≡ _ y≡z x≡y = trans x≡y y≡z - - step-≡˘ : ∀ (x {y z} : A) → y ≡ z → y ≡ x → x ≡ z - step-≡˘ _ y≡z y≡x = trans (sym y≡x) y≡z - - _∎ : ∀ (x : A) → x ≡ x - _∎ _ = refl - - syntax step-≡ x y≡z x≡y = x ≡⟨ x≡y ⟩ y≡z - syntax step-≡˘ x y≡z y≡x = x ≡˘⟨ y≡x ⟩ y≡z diff --git a/src/Relation/Binary/PropositionalEquality/Properties.agda b/src/Relation/Binary/PropositionalEquality/Properties.agda index 9515b77700..a420e45b64 100644 --- a/src/Relation/Binary/PropositionalEquality/Properties.agda +++ b/src/Relation/Binary/PropositionalEquality/Properties.agda @@ -23,6 +23,8 @@ open import Relation.Binary.Definitions import Relation.Binary.Properties.Setoid as Setoid open import Relation.Binary.PropositionalEquality.Core open import Relation.Unary using (Pred) +open import Relation.Binary.Reasoning.Syntax + private variable @@ -188,3 +190,16 @@ preorder A = Setoid.≈-preorder (setoid A) poset : Set a → Poset _ _ _ poset A = Setoid.≈-poset (setoid A) + +------------------------------------------------------------------------ +-- Reasoning + +-- This is a special instance of `Relation.Binary.Reasoning.Setoid`. +-- Rather than instantiating the latter with (setoid A), we reimplement +-- equation chains from scratch since then goals are printed much more +-- readably. +module ≡-Reasoning {a} {A : Set a} where + + open begin-syntax {A = A} _≡_ id public + open ≡-syntax {A = A} _≡_ trans public + open end-syntax {A = A} _≡_ refl public diff --git a/src/Relation/Binary/Reasoning/Syntax.agda b/src/Relation/Binary/Reasoning/Syntax.agda index 4837fe33f4..e8286bee30 100644 --- a/src/Relation/Binary/Reasoning/Syntax.agda +++ b/src/Relation/Binary/Reasoning/Syntax.agda @@ -22,7 +22,6 @@ open import Relation.Binary.PropositionalEquality.Core as P -- Relation/Binary/HeterogeneousEquality -- Effect/Monad/Partiality -- Effect/Monad/Partiality/All --- Relation/Binary/PropositionalEquality/Core -- Codata/Guarded/Stream/Relation/Binary/Pointwise -- Codata/Sized/Stream/Bisimilarity -- Function/Reasoning From cc0749cceef799f28c168d3c00a6f894972f27e0 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Wed, 18 Oct 2023 22:11:37 +0900 Subject: [PATCH 33/57] Add consistent deprecation warnings to old function hierarchy (#2163) --- CHANGELOG.md | 1 + .../Product/Function/Dependent/Setoid.agda | 9 ++- src/Function/Bijection.agda | 27 ++++++++ src/Function/Construct/Constant.agda | 64 +++++++++++++++++++ src/Function/Equality.agda | 21 ++++++ src/Function/Equivalence.agda | 35 ++++++++++ src/Function/Injection.agda | 31 +++++++-- src/Function/Inverse.agda | 35 ++++++++++ src/Function/LeftInverse.agda | 34 ++++++++++ src/Function/Properties/Equivalence.agda | 40 ++++++++---- src/Function/Properties/Injection.agda | 22 +++++-- src/Function/Properties/Surjection.agda | 22 +++++-- src/Function/Surjection.agda | 30 +++++++++ 13 files changed, 340 insertions(+), 31 deletions(-) create mode 100644 src/Function/Construct/Constant.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index 0ba8b736c4..9670c2085f 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1993,6 +1993,7 @@ New modules Function.Properties.Bijection Function.Properties.RightInverse Function.Properties.Surjection + Function.Construct.Constant ``` * The 'no infinite descent' principle for (accessible elements of) well-founded relations: diff --git a/src/Data/Product/Function/Dependent/Setoid.agda b/src/Data/Product/Function/Dependent/Setoid.agda index 9866971920..f290c7341b 100644 --- a/src/Data/Product/Function/Dependent/Setoid.agda +++ b/src/Data/Product/Function/Dependent/Setoid.agda @@ -16,11 +16,10 @@ open import Data.Product.Relation.Binary.Pointwise.Dependent as Σ open import Level using (Level) open import Function open import Function.Consequences -open import Function.Properties.Injection -open import Function.Properties.Surjection -open import Function.Properties.Equivalence -open import Function.Properties.RightInverse -import Function.Properties.Inverse as InverseProperties +open import Function.Properties.Injection using (mkInjection) +open import Function.Properties.Surjection using (mkSurjection; ↠⇒⇔) +open import Function.Properties.Equivalence using (mkEquivalence; ⇔⇒⟶; ⇔⇒⟵) +open import Function.Properties.RightInverse using (mkRightInverse) open import Relation.Binary.Core using (_=[_]⇒_) open import Relation.Binary.Bundles as B open import Relation.Binary.Indexed.Heterogeneous diff --git a/src/Function/Bijection.agda b/src/Function/Bijection.agda index bb7f2d2cce..ac32159e7a 100644 --- a/src/Function/Bijection.agda +++ b/src/Function/Bijection.agda @@ -38,6 +38,10 @@ record Bijective {f₁ f₂ t₁ t₂} left-inverse-of : from LeftInverseOf to left-inverse-of x = injective (right-inverse-of (to ⟨$⟩ x)) +{-# WARNING_ON_USAGE Bijective +"Warning: Bijective was deprecated in v2.0. +Please use Function.(Structures.)IsBijection instead." +#-} ------------------------------------------------------------------------ -- The set of all bijections between two setoids. @@ -74,6 +78,10 @@ record Bijection {f₁ f₂ t₁ t₂} } open LeftInverse left-inverse public using (to-from) +{-# WARNING_ON_USAGE Bijection +"Warning: Bijection was deprecated in v2.0. +Please use Function.(Bundles.)Bijection instead." +#-} ------------------------------------------------------------------------ -- The set of all bijections between two sets (i.e. bijections with @@ -83,6 +91,10 @@ infix 3 _⤖_ _⤖_ : ∀ {f t} → Set f → Set t → Set _ From ⤖ To = Bijection (P.setoid From) (P.setoid To) +{-# WARNING_ON_USAGE _⤖_ +"Warning: _⤖_ was deprecated in v2.0. +Please use Function.(Bundles.)mk⤖ instead." +#-} bijection : ∀ {f t} {From : Set f} {To : Set t} → (to : From → To) (from : To → From) → @@ -99,6 +111,11 @@ bijection to from inj invʳ = record } } } +{-# WARNING_ON_USAGE bijection +"Warning: bijection was deprecated in v2.0. +Please use either Function.Properties.Bijection.trans or +Function.Construct.Composition.bijection instead." +#-} ------------------------------------------------------------------------ -- Identity and composition. (Note that these proofs are superfluous, @@ -112,6 +129,11 @@ id {S = S} = record ; surjective = Surjection.surjective (Surj.id {S = S}) } } +{-# WARNING_ON_USAGE id +"Warning: id was deprecated in v2.0. +Please use either Function.Properties.Bijection.refl or +Function.Construct.Identity.bijection instead." +#-} infixr 9 _∘_ @@ -125,3 +147,8 @@ f ∘ g = record ; surjective = Surjection.surjective (Surj._∘_ (surjection f) (surjection g)) } } where open Bijection +{-# WARNING_ON_USAGE _∘_ +"Warning: _∘_ was deprecated in v2.0. +Please use either Function.Properties.Bijection.trans or +Function.Construct.Composition.bijection instead." +#-} diff --git a/src/Function/Construct/Constant.agda b/src/Function/Construct/Constant.agda new file mode 100644 index 0000000000..2f43929a87 --- /dev/null +++ b/src/Function/Construct/Constant.agda @@ -0,0 +1,64 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- The constant function +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Function.Construct.Constant where + +open import Function.Base using (const) +open import Function.Bundles +import Function.Definitions as Definitions +import Function.Structures as Structures +open import Level using (Level) +open import Relation.Binary.Core using (Rel) +open import Relation.Binary.Bundles using (Setoid) +open import Relation.Binary.Structures as B hiding (IsEquivalence) + +private + variable + a b ℓ₁ ℓ₂ : Level + A B : Set a + +------------------------------------------------------------------------ +-- Properties + +module _ (_≈₁_ : Rel A ℓ₁) (_≈₂_ : Rel B ℓ₂) where + + open Definitions + + congruent : ∀ {b} → b ≈₂ b → Congruent _≈₁_ _≈₂_ (const b) + congruent refl _ = refl + +------------------------------------------------------------------------ +-- Structures + +module _ + {≈₁ : Rel A ℓ₁} {≈₂ : Rel B ℓ₂} + (isEq₁ : B.IsEquivalence ≈₁) + (isEq₂ : B.IsEquivalence ≈₂) where + + open Structures ≈₁ ≈₂ + open B.IsEquivalence + + isCongruent : ∀ b → IsCongruent (const b) + isCongruent b = record + { cong = congruent ≈₁ ≈₂ (refl isEq₂) + ; isEquivalence₁ = isEq₁ + ; isEquivalence₂ = isEq₂ + } + +------------------------------------------------------------------------ +-- Setoid bundles + +module _ (S : Setoid a ℓ₂) (T : Setoid b ℓ₂) where + + open Setoid + + function : Carrier T → Func S T + function b = record + { to = const b + ; cong = congruent (_≈_ S) (_≈_ T) (refl T) + } diff --git a/src/Function/Equality.agda b/src/Function/Equality.agda index 4ce068b45c..7451aaf3f7 100644 --- a/src/Function/Equality.agda +++ b/src/Function/Equality.agda @@ -5,6 +5,7 @@ ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --warn=noUserWarning #-} module Function.Equality where @@ -34,6 +35,10 @@ record Π {f₁ f₂ t₁ t₂} field _⟨$⟩_ : (x : Setoid.Carrier From) → IndexedSetoid.Carrier To x cong : Setoid._≈_ From =[ _⟨$⟩_ ]⇒ IndexedSetoid._≈_ To +{-# WARNING_ON_USAGE Π +"Warning: Π was deprecated in v2.0. +Please use Function.Dependent.Bundles.Func instead." +#-} open Π public @@ -41,12 +46,20 @@ infixr 0 _⟶_ _⟶_ : ∀ {f₁ f₂ t₁ t₂} → Setoid f₁ f₂ → Setoid t₁ t₂ → Set _ From ⟶ To = Π From (Trivial.indexedSetoid To) +{-# WARNING_ON_USAGE _⟶_ +"Warning: _⟶_ was deprecated in v2.0. +Please use Function.(Bundles.)Func instead." +#-} ------------------------------------------------------------------------ -- Identity and composition. id : ∀ {a₁ a₂} {A : Setoid a₁ a₂} → A ⟶ A id = record { _⟨$⟩_ = Fun.id; cong = Fun.id } +{-# WARNING_ON_USAGE id +"Warning: id was deprecated in v2.0. +Please use Function.Construct.Identity.function instead." +#-} infixr 9 _∘_ @@ -58,6 +71,10 @@ f ∘ g = record { _⟨$⟩_ = Fun._∘_ (_⟨$⟩_ f) (_⟨$⟩_ g) ; cong = Fun._∘_ (cong f) (cong g) } +{-# WARNING_ON_USAGE _∘_ +"Warning: _∘_ was deprecated in v2.0. +Please use Function.Construct.Composition.function instead." +#-} -- Constant equality-preserving function. @@ -68,6 +85,10 @@ const {B = B} b = record { _⟨$⟩_ = Fun.const b ; cong = Fun.const (Setoid.refl B) } +{-# WARNING_ON_USAGE const +"Warning: const was deprecated in v2.0. +Please use Function.Construct.Constant.function instead." +#-} ------------------------------------------------------------------------ -- Function setoids diff --git a/src/Function/Equivalence.agda b/src/Function/Equivalence.agda index b25f2a062f..1e3652da7f 100644 --- a/src/Function/Equivalence.agda +++ b/src/Function/Equivalence.agda @@ -31,6 +31,10 @@ record Equivalence {f₁ f₂ t₁ t₂} field to : From ⟶ To from : To ⟶ From +{-# WARNING_ON_USAGE Equivalence +"Warning: Equivalence was deprecated in v2.0. +Please use Function.(Bundles.)Equivalence instead." +#-} ------------------------------------------------------------------------ -- The set of all equivalences between two sets (i.e. equivalences @@ -40,6 +44,10 @@ infix 3 _⇔_ _⇔_ : ∀ {f t} → Set f → Set t → Set _ From ⇔ To = Equivalence (P.setoid From) (P.setoid To) +{-# WARNING_ON_USAGE _⇔_ +"Warning: _⇔_ was deprecated in v2.0. +Please use Function.(Bundles.)_⇔_ instead." +#-} equivalence : ∀ {f t} {From : Set f} {To : Set t} → (From → To) → (To → From) → From ⇔ To @@ -47,6 +55,10 @@ equivalence to from = record { to = →-to-⟶ to ; from = →-to-⟶ from } +{-# WARNING_ON_USAGE equivalence +"Warning: equivalence was deprecated in v2.0. +Please use Function.Properties.Equivalence.mkEquivalence instead." +#-} ------------------------------------------------------------------------ -- Equivalence is an equivalence relation @@ -58,6 +70,11 @@ id {x = S} = record { to = F.id ; from = F.id } +{-# WARNING_ON_USAGE id +"Warning: id was deprecated in v2.0. +Please use Function.Properties.Equivalence.refl or +Function.Construct.Identity.equivalence instead." +#-} infixr 9 _∘_ @@ -69,6 +86,11 @@ f ∘ g = record { to = to f ⟪∘⟫ to g ; from = from g ⟪∘⟫ from f } where open Equivalence +{-# WARNING_ON_USAGE _∘_ +"Warning: _∘_ was deprecated in v2.0. +Please use Function.Properties.Equivalence.trans or +Function.Construct.Composition.equivalence instead." +#-} -- Symmetry. @@ -79,6 +101,11 @@ sym eq = record { from = to ; to = from } where open Equivalence eq +{-# WARNING_ON_USAGE sym +"Warning: sym was deprecated in v2.0. +Please use Function.Properties.Equivalence.sym or +Function.Construct.Symmetry.equivalence instead." +#-} -- For fixed universe levels we can construct setoids. @@ -92,6 +119,10 @@ setoid s₁ s₂ = record ; trans = flip _∘_ } } +{-# WARNING_ON_USAGE setoid +"Warning: setoid was deprecated in v2.0. +Please use Function.Properties.Equivalence.setoid instead." +#-} ⇔-setoid : (ℓ : Level) → Setoid (suc ℓ) ℓ ⇔-setoid ℓ = record @@ -103,6 +134,10 @@ setoid s₁ s₂ = record ; trans = flip _∘_ } } +{-# WARNING_ON_USAGE ⇔-setoid +"Warning: ⇔-setoid was deprecated in v2.0. +Please use Function.Properties.Equivalence.⇔-setoid instead." +#-} ------------------------------------------------------------------------ -- Transformations diff --git a/src/Function/Injection.agda b/src/Function/Injection.agda index 910c91aa35..462b8b3428 100644 --- a/src/Function/Injection.agda +++ b/src/Function/Injection.agda @@ -7,11 +7,6 @@ {-# OPTIONS --cubical-compatible --safe #-} {-# OPTIONS --warn=noUserWarning #-} --- Note: use of the standard function hierarchy is encouraged. The --- module `Function` re-exports `Injective`, `IsInjection` and --- `Injection`. The alternative definitions found in this file will --- eventually be deprecated. - module Function.Injection where {-# WARNING_ON_IMPORT @@ -35,6 +30,10 @@ Injective {A = A} {B} f = ∀ {x y} → f ⟨$⟩ x ≈₂ f ⟨$⟩ y → x ≈ where open Setoid A renaming (_≈_ to _≈₁_) open Setoid B renaming (_≈_ to _≈₂_) +{-# WARNING_ON_USAGE Injective +"Warning: Injective was deprecated in v2.0. +Please use Function.(Definitions.)Injective instead." +#-} ------------------------------------------------------------------------ -- The set of all injections between two setoids @@ -47,6 +46,10 @@ record Injection {f₁ f₂ t₁ t₂} injective : Injective to open Π to public +{-# WARNING_ON_USAGE Injection +"Warning: Injection was deprecated in v2.0. +Please use Function.(Bundles.)Injection instead." +#-} ------------------------------------------------------------------------ -- The set of all injections from one set to another (i.e. injections @@ -56,6 +59,10 @@ infix 3 _↣_ _↣_ : ∀ {f t} → Set f → Set t → Set _ From ↣ To = Injection (P.setoid From) (P.setoid To) +{-# WARNING_ON_USAGE _↣_ +"Warning: _↣_ was deprecated in v2.0. +Please use Function.(Bundles.)_↣_ instead." +#-} injection : ∀ {f t} {From : Set f} {To : Set t} → (to : From → To) → (∀ {x y} → to x ≡ to y → x ≡ y) → From ↣ To @@ -66,6 +73,10 @@ injection to injective = record } ; injective = injective } +{-# WARNING_ON_USAGE injection +"Warning: injection was deprecated in v2.0. +Please use Function.(Bundles.)mk↣ instead." +#-} ------------------------------------------------------------------------ -- Identity and composition. @@ -77,6 +88,11 @@ id = record { to = F.id ; injective = Fun.id } +{-# WARNING_ON_USAGE id +"Warning: id was deprecated in v2.0. +Please use Function.Properties.Injection.refl or +Function.Construct.Identity.injection instead." +#-} _∘_ : ∀ {f₁ f₂ m₁ m₂ t₁ t₂} {F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂} → @@ -85,3 +101,8 @@ f ∘ g = record { to = to f ⟪∘⟫ to g ; injective = (λ {_} → injective g) ⟨∘⟩ injective f } where open Injection +{-# WARNING_ON_USAGE _∘_ +"Warning: _∘_ was deprecated in v2.0. +Please use Function.Properties.Injection.trans or +Function.Construct.Composition.injection instead." +#-} diff --git a/src/Function/Inverse.agda b/src/Function/Inverse.agda index e2cb395ea5..e9d796eddf 100644 --- a/src/Function/Inverse.agda +++ b/src/Function/Inverse.agda @@ -35,6 +35,10 @@ record _InverseOf_ {f₁ f₂ t₁ t₂} field left-inverse-of : from LeftInverseOf to right-inverse-of : from RightInverseOf to +{-# WARNING_ON_USAGE _InverseOf_ +"Warning: _InverseOf_ was deprecated in v2.0. +Please use Function.(Structures.)IsInverse instead." +#-} ------------------------------------------------------------------------ -- The set of all inverses between two setoids @@ -74,6 +78,10 @@ record Inverse {f₁ f₂ t₁ t₂} open Bijection bijection public using (equivalence; surjective; surjection; right-inverse; to-from; from-to) +{-# WARNING_ON_USAGE Inverse +"Warning: Inverse was deprecated in v2.0. +Please use Function.(Bundles.)Inverse instead." +#-} ------------------------------------------------------------------------ -- The set of all inverses between two sets (i.e. inverses with @@ -83,9 +91,17 @@ infix 3 _↔_ _↔̇_ _↔_ : ∀ {f t} → Set f → Set t → Set _ From ↔ To = Inverse (P.setoid From) (P.setoid To) +{-# WARNING_ON_USAGE _↔_ +"Warning: _↔_ was deprecated in v2.0. +Please use Function.(Bundles.)_↔_ instead." +#-} _↔̇_ : ∀ {i f t} {I : Set i} → Pred I f → Pred I t → Set _ From ↔̇ To = ∀ {i} → From i ↔ To i +{-# WARNING_ON_USAGE _↔̇_ +"Warning: _↔̇_ was deprecated in v2.0. +Please use Function.Indexed.(Bundles.)_↔ᵢ_ instead." +#-} inverse : ∀ {f t} {From : Set f} {To : Set t} → (to : From → To) (from : To → From) → @@ -100,6 +116,10 @@ inverse to from from∘to to∘from = record ; right-inverse-of = to∘from } } +{-# WARNING_ON_USAGE inverse +"Warning: inverse was deprecated in v2.0. +Please use Function.(Bundles.)mk↔ instead." +#-} ------------------------------------------------------------------------ -- If two setoids are in bijective correspondence, then there is an @@ -131,6 +151,11 @@ id {x = S} = record ; right-inverse-of = LeftInverse.left-inverse-of id′ } } where id′ = Left.id {S = S} +{-# WARNING_ON_USAGE id +"Warning: id was deprecated in v2.0. +Please use either Function.Properties.Inverse.refl or +Function.Construct.Identity.inverse instead." +#-} -- Transitivity @@ -148,6 +173,11 @@ f ∘ g = record ; right-inverse-of = LeftInverse.left-inverse-of (Left._∘_ (right-inverse g) (right-inverse f)) } } where open Inverse +{-# WARNING_ON_USAGE _∘_ +"Warning: _∘_ was deprecated in v2.0. +Please use either Function.Properties.Inverse.trans or +Function.Construct.Composition.inverse instead." +#-} -- Symmetry. @@ -161,6 +191,11 @@ sym inv = record ; right-inverse-of = left-inverse-of } } where open Inverse inv +{-# WARNING_ON_USAGE sym +"Warning: sym was deprecated in v2.0. +Please use either Function.Properties.Inverse.sym or +Function.Construct.Symmetry.inverse instead." +#-} ------------------------------------------------------------------------ -- Transformations diff --git a/src/Function/LeftInverse.agda b/src/Function/LeftInverse.agda index 7ff0fda218..3824ca95d3 100644 --- a/src/Function/LeftInverse.agda +++ b/src/Function/LeftInverse.agda @@ -31,11 +31,19 @@ _LeftInverseOf_ : To ⟶ From → From ⟶ To → Set _ _LeftInverseOf_ {From = From} f g = ∀ x → f ⟨$⟩ (g ⟨$⟩ x) ≈ x where open Setoid From +{-# WARNING_ON_USAGE _LeftInverseOf_ +"Warning: _LeftInverseOf_ was deprecated in v2.0. +Please use Function.(Structures.)IsRightInverse instead." +#-} _RightInverseOf_ : ∀ {f₁ f₂ t₁ t₂} {From : Setoid f₁ f₂} {To : Setoid t₁ t₂} → To ⟶ From → From ⟶ To → Set _ f RightInverseOf g = g LeftInverseOf f +{-# WARNING_ON_USAGE _RightInverseOf_ +"Warning: _RightInverseOf_ was deprecated in v2.0. +Please use Function.(Structures.)IsLeftInverse instead." +#-} ------------------------------------------------------------------------ -- The set of all left inverses between two setoids. @@ -74,12 +82,20 @@ record LeftInverse {f₁ f₂ t₁ t₂} from ⟨$⟩ y ≈⟨ Eq.cong from (T.sym to-x≈y) ⟩ from ⟨$⟩ (to ⟨$⟩ x) ≈⟨ left-inverse-of x ⟩ x ∎ +{-# WARNING_ON_USAGE LeftInverse +"Warning: LeftInverse was deprecated in v2.0. +Please use Function.(Bundles.)RightInverse instead." +#-} -- The set of all right inverses between two setoids. RightInverse : ∀ {f₁ f₂ t₁ t₂} (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) → Set _ RightInverse From To = LeftInverse To From +{-# WARNING_ON_USAGE RightInverse +"Warning: RightInverse was deprecated in v2.0. +Please use Function.(Bundles.)LeftInverse instead." +#-} ------------------------------------------------------------------------ -- The set of all left inverses from one set to another (i.e. left @@ -91,6 +107,10 @@ infix 3 _↞_ _↞_ : ∀ {f t} → Set f → Set t → Set _ From ↞ To = LeftInverse (P.setoid From) (P.setoid To) +{-# WARNING_ON_USAGE _↞_ +"Warning: _↞_ was deprecated in v2.0. +Please use Function.(Bundles.)_↪_ instead." +#-} leftInverse : ∀ {f t} {From : Set f} {To : Set t} → (to : From → To) (from : To → From) → @@ -101,6 +121,10 @@ leftInverse to from invˡ = record ; from = Eq.→-to-⟶ from ; left-inverse-of = invˡ } +{-# WARNING_ON_USAGE leftInverse +"Warning: leftInverse was deprecated in v2.0. +Please use Function.(Bundles.)mk↪ instead." +#-} ------------------------------------------------------------------------ -- Identity and composition. @@ -111,6 +135,11 @@ id {S = S} = record ; from = Eq.id ; left-inverse-of = λ _ → Setoid.refl S } +{-# WARNING_ON_USAGE id +"Warning: id was deprecated in v2.0. +Please use either Function.Properties.RightInverse.refl or +Function.Construct.Identity.rightInverse instead." +#-} infixr 9 _∘_ @@ -128,3 +157,8 @@ _∘_ {F = F} f g = record where open LeftInverse open EqReasoning F +{-# WARNING_ON_USAGE _∘_ +"Warning: _∘_ was deprecated in v2.0. +Please use either Function.Properties.RightInverse.trans or +Function.Construct.Composition.rightInverse instead." +#-} diff --git a/src/Function/Properties/Equivalence.agda b/src/Function/Properties/Equivalence.agda index 308a29525d..ad1264ec7c 100644 --- a/src/Function/Properties/Equivalence.agda +++ b/src/Function/Properties/Equivalence.agda @@ -11,6 +11,7 @@ module Function.Properties.Equivalence where open import Function.Bundles open import Level +open import Relation.Binary.Definitions open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) import Relation.Binary.PropositionalEquality.Properties as Eq @@ -21,7 +22,7 @@ import Function.Construct.Composition as Composition private variable - a ℓ : Level + a b c ℓ ℓ₁ ℓ₂ ℓ₃ : Level A B : Set a S T : Setoid a ℓ @@ -36,13 +37,37 @@ mkEquivalence f g = record ; from-cong = cong g } where open Func +⟶×⟵⇒⇔ : A ⟶ B → B ⟶ A → A ⇔ B +⟶×⟵⇒⇔ = mkEquivalence + +------------------------------------------------------------------------ +-- Destructors + +⇔⇒⟶ : A ⇔ B → A ⟶ B +⇔⇒⟶ = Equivalence.toFunction + +⇔⇒⟵ : A ⇔ B → B ⟶ A +⇔⇒⟵ = Equivalence.fromFunction + ------------------------------------------------------------------------ -- Setoid bundles +refl : Reflexive (Equivalence {a} {ℓ}) +refl {x = x} = Identity.equivalence x + +sym : Sym (Equivalence {a} {ℓ₁} {b} {ℓ₂}) + (Equivalence {b} {ℓ₂} {a} {ℓ₁}) +sym = Symmetry.equivalence + +trans : Trans (Equivalence {a} {ℓ₁} {b} {ℓ₂}) + (Equivalence {b} {ℓ₂} {c} {ℓ₃}) + (Equivalence {a} {ℓ₁} {c} {ℓ₃}) +trans = Composition.equivalence + isEquivalence : IsEquivalence (Equivalence {a} {ℓ}) isEquivalence = record - { refl = λ {x} → Identity.equivalence x - ; sym = Symmetry.equivalence + { refl = refl + ; sym = sym ; trans = Composition.equivalence } @@ -69,12 +94,3 @@ setoid s₁ s₂ = record ; _≈_ = _⇔_ ; isEquivalence = ⇔-isEquivalence } - -⟶×⟵⇒⇔ : A ⟶ B → B ⟶ A → A ⇔ B -⟶×⟵⇒⇔ = mkEquivalence - -⇔⇒⟶ : A ⇔ B → A ⟶ B -⇔⇒⟶ = Equivalence.toFunction - -⇔⇒⟵ : A ⇔ B → B ⟶ A -⇔⇒⟵ = Equivalence.fromFunction diff --git a/src/Function/Properties/Injection.agda b/src/Function/Properties/Injection.agda index bca8ac4dc5..0e773cf9b8 100644 --- a/src/Function/Properties/Injection.agda +++ b/src/Function/Properties/Injection.agda @@ -15,13 +15,14 @@ import Function.Construct.Identity as Identity import Function.Construct.Composition as Compose open import Level using (Level) open import Data.Product.Base using (proj₁; proj₂) -open import Relation.Binary.PropositionalEquality +open import Relation.Binary.Definitions +open import Relation.Binary.PropositionalEquality using () open import Relation.Binary using (Setoid) import Relation.Binary.Reasoning.Setoid as SetoidReasoning private variable - a b ℓ : Level + a b c ℓ ℓ₁ ℓ₂ ℓ₃ : Level A B : Set a T S U : Setoid a ℓ @@ -43,10 +44,21 @@ mkInjection f injective = record ↣⇒⟶ = Injection.function ------------------------------------------------------------------------ --- Properties +-- Setoid properties + +refl : Reflexive (Injection {a} {ℓ}) +refl {x = x} = Identity.injection x + +trans : Trans (Injection {a} {ℓ₁} {b} {ℓ₂}) + (Injection {b} {ℓ₂} {c} {ℓ₃}) + (Injection {a} {ℓ₁} {c} {ℓ₃}) +trans = Compose.injection + +------------------------------------------------------------------------ +-- Propositonal properties ↣-refl : Injection S S -↣-refl = Identity.injection _ +↣-refl = refl ↣-trans : Injection S T → Injection T U → Injection S U -↣-trans = Compose.injection +↣-trans = trans diff --git a/src/Function/Properties/Surjection.agda b/src/Function/Properties/Surjection.agda index d16670639c..5c2b235506 100644 --- a/src/Function/Properties/Surjection.agda +++ b/src/Function/Properties/Surjection.agda @@ -11,15 +11,18 @@ module Function.Properties.Surjection where open import Function.Base open import Function.Definitions open import Function.Bundles +import Function.Construct.Identity as Identity +import Function.Construct.Composition as Compose open import Level using (Level) open import Data.Product.Base using (proj₁; proj₂) -open import Relation.Binary.PropositionalEquality -open import Relation.Binary using (Setoid) +import Relation.Binary.PropositionalEquality as P +open import Relation.Binary.Definitions +open import Relation.Binary.Bundles using (Setoid) import Relation.Binary.Reasoning.Setoid as SetoidReasoning private variable - a b ℓ : Level + a b c ℓ ℓ₁ ℓ₂ ℓ₃ : Level A B : Set a T S : Setoid a ℓ @@ -41,13 +44,24 @@ mkSurjection f surjective = record ↠⇒⟶ = Surjection.function ↠⇒↪ : A ↠ B → B ↪ A -↠⇒↪ s = mk↪ {from = to} λ { refl → proj₂ (strictlySurjective _)} +↠⇒↪ s = mk↪ {from = to} λ { P.refl → proj₂ (strictlySurjective _)} where open Surjection s ↠⇒⇔ : A ↠ B → A ⇔ B ↠⇒⇔ s = mk⇔ to (proj₁ ∘ surjective) where open Surjection s +------------------------------------------------------------------------ +-- Setoid properties + +refl : Reflexive (Surjection {a} {ℓ}) +refl {x = x} = Identity.surjection x + +trans : Trans (Surjection {a} {ℓ₁} {b} {ℓ₂}) + (Surjection {b} {ℓ₂} {c} {ℓ₃}) + (Surjection {a} {ℓ₁} {c} {ℓ₃}) +trans = Compose.surjection + ------------------------------------------------------------------------ -- Other diff --git a/src/Function/Surjection.agda b/src/Function/Surjection.agda index bd96c50efb..2fdfefa9af 100644 --- a/src/Function/Surjection.agda +++ b/src/Function/Surjection.agda @@ -33,6 +33,10 @@ record Surjective {f₁ f₂ t₁ t₂} field from : To ⟶ From right-inverse-of : from RightInverseOf to +{-# WARNING_ON_USAGE Surjective +"Warning: Surjective was deprecated in v2.0. +Please use Function.(Definitions.)Surjective instead." +#-} ------------------------------------------------------------------------ -- The set of all surjections from one setoid to another. @@ -67,6 +71,10 @@ record Surjection {f₁ f₂ t₁ t₂} { to = to ; from = from } +{-# WARNING_ON_USAGE Surjection +"Warning: Surjection was deprecated in v2.0. +Please use Function.(Bundles.)Surjection instead." +#-} -- Right inverses can be turned into surjections. @@ -80,6 +88,10 @@ fromRightInverse r = record ; right-inverse-of = left-inverse-of } } where open LeftInverse r +{-# WARNING_ON_USAGE fromRightInverse +"Warning: fromRightInverse was deprecated in v2.0. +Please use Function.(Properties.)RightInverse.RightInverse⇒Surjection instead." +#-} ------------------------------------------------------------------------ -- The set of all surjections from one set to another (i.e. sujections @@ -89,6 +101,10 @@ infix 3 _↠_ _↠_ : ∀ {f t} → Set f → Set t → Set _ From ↠ To = Surjection (P.setoid From) (P.setoid To) +{-# WARNING_ON_USAGE _↠_ +"Warning: _↠_ was deprecated in v2.0. +Please use Function.(Bundles.)_↠_ instead." +#-} surjection : ∀ {f t} {From : Set f} {To : Set t} → (to : From → To) (from : To → From) → @@ -101,6 +117,10 @@ surjection to from surjective = record ; right-inverse-of = surjective } } +{-# WARNING_ON_USAGE surjection +"Warning: surjection was deprecated in v2.0. +Please use Function.(Bundles.)mk↠ instead." +#-} ------------------------------------------------------------------------ -- Identity and composition. @@ -113,6 +133,11 @@ id {S = S} = record ; right-inverse-of = LeftInverse.left-inverse-of id′ } } where id′ = Left.id {S = S} +{-# WARNING_ON_USAGE id +"Warning: id was deprecated in v2.0. +Please use Function.Properties.Surjection.refl or +Function.Construct.Identity.surjection instead." +#-} infixr 9 _∘_ @@ -129,3 +154,8 @@ f ∘ g = record where open Surjection g∘f = Left._∘_ (right-inverse g) (right-inverse f) +{-# WARNING_ON_USAGE _∘_ +"Warning: _∘_ was deprecated in v2.0. +Please use Function.Properties.Surjection.trans or +Function.Construct.Composition.surjection instead." +#-} From fe5fddc659ad895d0085f19d678973f9cec2f250 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Thu, 19 Oct 2023 09:40:49 +0900 Subject: [PATCH 34/57] Rename symmetric reasoning combinators. Minimal set of changes (#2160) --- CHANGELOG.md | 63 ++++++-- README/Data/Fin/Relation/Unary/Top.agda | 4 +- .../Vec/Relation/Binary/Equality/Cast.agda | 24 +-- README/Tactic/Cong.agda | 6 +- .../Properties/HeytingCommutativeRing.agda | 16 +- src/Algebra/Consequences/Setoid.agda | 10 +- src/Algebra/Construct/LexProduct/Inner.agda | 36 ++--- src/Algebra/Construct/LiftedChoice.agda | 6 +- .../Construct/NaturalChoice/MinMaxOp.agda | 16 +- .../Construct/NaturalChoice/MinOp.agda | 22 +-- .../Lattice/Morphism/LatticeMonomorphism.agda | 4 +- .../Lattice/Properties/BooleanAlgebra.agda | 96 ++++++------ src/Algebra/Lattice/Properties/Lattice.agda | 4 +- src/Algebra/Module/Consequences.agda | 8 +- src/Algebra/Morphism/Consequences.agda | 8 +- src/Algebra/Morphism/GroupMonomorphism.agda | 6 +- src/Algebra/Morphism/MagmaMonomorphism.agda | 12 +- src/Algebra/Morphism/MonoidMonomorphism.agda | 4 +- src/Algebra/Morphism/RingMonomorphism.agda | 12 +- src/Algebra/Properties/AbelianGroup.agda | 2 +- .../CommutativeMonoid/Mult/TCOptimised.agda | 2 +- .../Properties/CommutativeMonoid/Sum.agda | 4 +- .../Properties/CommutativeSemigroup.agda | 4 +- src/Algebra/Properties/Group.agda | 16 +- src/Algebra/Properties/Monoid/Mult.agda | 2 +- .../Properties/Monoid/Mult/TCOptimised.agda | 8 +- src/Algebra/Properties/Monoid/Sum.agda | 4 +- src/Algebra/Properties/Semiring/Binomial.agda | 34 ++--- src/Algebra/Properties/Semiring/Exp.agda | 8 +- .../Semiring/Exp/TailRecursiveOptimised.agda | 2 +- src/Algebra/Properties/Semiring/Mult.agda | 4 +- .../Properties/Semiring/Mult/TCOptimised.agda | 2 +- .../Solver/Ring/NaturalCoefficients.agda | 2 +- .../Stream/Relation/Binary/Pointwise.agda | 35 ++++- src/Codata/Sized/Colist/Properties.agda | 4 +- src/Data/Fin/Properties.agda | 16 +- src/Data/Fin/Subset/Properties.agda | 2 +- src/Data/Integer/DivMod.agda | 2 +- src/Data/Integer/Properties.agda | 102 ++++++------- src/Data/List/Effectful.agda | 14 +- src/Data/List/Properties.agda | 4 +- .../Relation/Binary/BagAndSetEquality.agda | 8 +- .../Binary/Permutation/Propositional.agda | 6 +- .../Permutation/Propositional/Properties.agda | 10 +- .../Relation/Binary/Permutation/Setoid.agda | 8 +- .../Binary/Permutation/Setoid/Properties.agda | 20 +-- .../Binary/Subset/Setoid/Properties.agda | 2 +- .../List/Relation/Unary/Any/Properties.agda | 4 +- src/Data/Nat/Binary/Properties.agda | 8 +- src/Data/Nat/Combinatorics.agda | 30 ++-- src/Data/Nat/Combinatorics/Specification.agda | 8 +- src/Data/Nat/DivMod.agda | 50 +++---- src/Data/Nat/DivMod/WithK.agda | 2 +- src/Data/Nat/Divisibility.agda | 6 +- src/Data/Nat/LCM.agda | 6 +- src/Data/Nat/Primality.agda | 6 +- src/Data/Nat/Properties.agda | 18 +-- src/Data/Rational/Properties.agda | 53 +++---- .../Rational/Unnormalised/Properties.agda | 117 +++++++-------- src/Data/Tree/Binary/Zipper/Properties.agda | 12 +- src/Data/Tree/Rose/Properties.agda | 4 +- src/Data/Vec/Functional/Properties.agda | 8 +- src/Data/Vec/Properties.agda | 30 ++-- .../Vec/Relation/Binary/Equality/Cast.agda | 53 ++++--- src/Function/Properties/Surjection.agda | 2 +- .../ReflexiveTransitive/Properties.agda | 2 +- .../Binary/Construct/NaturalOrder/Right.agda | 2 +- .../Binary/HeterogeneousEquality.agda | 14 +- .../Binary/Reasoning/PartialOrder.agda | 4 +- src/Relation/Binary/Reasoning/Preorder.agda | 2 +- src/Relation/Binary/Reasoning/Setoid.agda | 6 +- .../Binary/Reasoning/StrictPartialOrder.agda | 4 +- src/Relation/Binary/Reasoning/Syntax.agda | 137 ++++++++++++++---- src/Tactic/Cong.agda | 2 +- src/Tactic/MonoidSolver.agda | 4 +- .../Core/Polynomial/Homomorphism/Lemmas.agda | 8 +- src/Text/Pretty/Core.agda | 2 +- 77 files changed, 716 insertions(+), 572 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 9670c2085f..4d6a8c2ccd 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -849,6 +849,7 @@ Non-backwards compatible changes IO.Effectful IO.Instances ``` + ### (Issue #2096) Introduction of flipped relation symbol for `Relation.Binary.Bundles.Preorder` * Previously, the relation symbol `_∼_` was (notationally) symmetric, so that its @@ -989,6 +990,53 @@ Non-backwards compatible changes Relation.Binary.Reasoning.PartialOrder ``` +### More modular design of equational reasoning. + +* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports + a range of modules containing pre-existing reasoning combinator syntax. + +* This makes it possible to add new or rename existing reasoning combinators to a + pre-existing `Reasoning` module in just a couple of lines + (e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`) + +* One pre-requisite for that is that `≡-Reasoning` has been moved from + `Relation.Binary.PropositionalEquality.Core` (which shouldn't be + imported anyway as it's a `Core` module) to + `Relation.Binary.PropositionalEquality.Properties`. + It is still exported by `Relation.Binary.PropositionalEquality`. + +### Renaming of symmetric combinators for equational reasoning + +* We've had various complaints about the symmetric version of reasoning combinators + that use the syntax `_R˘⟨_⟩_` for some relation `R`, (e.g. `_≡˘⟨_⟩_` and `_≃˘⟨_⟩_`) + introduced in `v1.0`. In particular: + 1. The symbol `˘` is hard to type. + 2. The symbol `˘` doesn't convey the direction of the equality + 3. The left brackets aren't vertically aligned with the left brackets of the non-symmetric version. + +* To address these problems we have renamed all the symmetric versions of the + combinators from `_R˘⟨_⟩_` to `_R⟨_⟨_` (the direction of the last bracket is flipped + to indicate the quality goes from right to left). + +* The old combinators still exist but have been deprecated. However due to + [Agda issue #5617](https://github.com/agda/agda/issues/5617), the deprecation warnings + don't fire correctly. We will not remove the old combinators before the above issue is + addressed. However, we still encourage migrating to the new combinators! + +* On a Linux-based system, the following command was used to globally migrate all uses of the + old combinators to the new ones in the standard library itself. + It *may* be useful when trying to migrate your own code: + ```bash + find . -type f -name "*.agda" -print0 | xargs -0 sed -i -e 's/˘⟨\(.*\)⟩/⟨\1⟨/g' + ``` + USUAL CAVEATS: It has not been widely tested and the standard library developers are not + responsible for the results of this command. It is strongly recommended you back up your + work before you attempt to run it. + +* NOTE: this refactoring may require some extra bracketing around the operator `_⟨_⟩_` from + `Function.Base` if the `_⟨_⟩_` operator is used within the reasoning proof. The symptom + for this is a `Don't know how to parse` error. + ### Other * In accordance with changes to the flags in Agda 2.6.3, all modules that previously used @@ -1228,21 +1276,6 @@ Major improvements * `RawLattice` has been moved from `Algebra.Lattice.Bundles` to this new module. * In `Relation.Binary.Reasoning.Base.Triple`, added a new parameter `<-irrefl : Irreflexive _≈_ _<_` - -### More modular design of equational reasoning. - -* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports - a range of modules containing pre-existing reasoning combinator syntax. - -* This makes it possible to add new or rename existing reasoning combinators to a - pre-existing `Reasoning` module in just a couple of lines - (e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`) - -* One pre-requisite for that is that `≡-Reasoning` has been moved from - `Relation.Binary.PropositionalEquality.Core` (which shouldn't be - imported anyway as it's a `Core` module) to - `Relation.Binary.PropositionalEquality.Properties`. - It is still exported by `Relation.Binary.PropositionalEquality`. Deprecated modules ------------------ diff --git a/README/Data/Fin/Relation/Unary/Top.agda b/README/Data/Fin/Relation/Unary/Top.agda index e602932827..390a48d58f 100644 --- a/README/Data/Fin/Relation/Unary/Top.agda +++ b/README/Data/Fin/Relation/Unary/Top.agda @@ -76,8 +76,8 @@ opposite-prop {suc n} i with view i ... | ‵fromℕ rewrite toℕ-fromℕ n | n∸n≡0 n = refl ... | ‵inject₁ j = begin suc (toℕ (opposite j)) ≡⟨ cong suc (opposite-prop j) ⟩ - suc (n ∸ suc (toℕ j)) ≡˘⟨ +-∸-assoc 1 (toℕ ⊖-monoʳ-≥-≤ (suc n) {suc m} {suc o} (s≤s m≤o) = begin suc n ⊖ suc m ≡⟨ [1+m]⊖[1+n]≡m⊖n n m ⟩ n ⊖ m ≤⟨ ⊖-monoʳ-≥-≤ n m≤o ⟩ - n ⊖ o ≡˘⟨ [1+m]⊖[1+n]≡m⊖n n o ⟩ + n ⊖ o ≡⟨ [1+m]⊖[1+n]≡m⊖n n o ⟨ suc n ⊖ suc o ∎ where open ≤-Reasoning ⊖-monoˡ-≤ : ∀ n → (_⊖ n) Preserves ℕ._≤_ ⟶ _≤_ @@ -759,12 +759,12 @@ sign-⊖-≰ = sign-⊖-< ∘ ℕ.≰⇒> ⊖-monoˡ-≤ (suc n) {_} {suc o} z≤n = begin zero ⊖ suc n ≤⟨ ⊖-monoʳ-≥-≤ 0 (ℕ.n≤1+n n) ⟩ zero ⊖ n ≤⟨ ⊖-monoˡ-≤ n z≤n ⟩ - o ⊖ n ≡˘⟨ [1+m]⊖[1+n]≡m⊖n o n ⟩ + o ⊖ n ≡⟨ [1+m]⊖[1+n]≡m⊖n o n ⟨ suc o ⊖ suc n ∎ where open ≤-Reasoning ⊖-monoˡ-≤ (suc n) {suc m} {suc o} (s≤s m≤o) = begin suc m ⊖ suc n ≡⟨ [1+m]⊖[1+n]≡m⊖n m n ⟩ m ⊖ n ≤⟨ ⊖-monoˡ-≤ n m≤o ⟩ - o ⊖ n ≡˘⟨ [1+m]⊖[1+n]≡m⊖n o n ⟩ + o ⊖ n ≡⟨ [1+m]⊖[1+n]≡m⊖n o n ⟨ suc o ⊖ suc n ∎ where open ≤-Reasoning ⊖-monoʳ->-< : ∀ p → (p ⊖_) Preserves ℕ._>_ ⟶ _<_ @@ -777,19 +777,19 @@ sign-⊖-≰ = sign-⊖-< ∘ ℕ.≰⇒> ⊖-monoʳ->-< (suc p) {suc m} {suc n} (s-< p m_ *-monoˡ-<-neg i@(-[1+ _ ]) {j} {k} j j *-cancelˡ-<-nonPos {i} {j} +0 (+<+ ()) *-cancelˡ-<-nonPos {i} {j} k@(-[1+ _ ]) ki>= pam as) unfold-⊛ fs as = begin fs ⊛ as - ≡˘⟨ concatMap-cong (λ f → P.cong (map f) (concatMap-pure as)) fs ⟩ + ≡⟨ concatMap-cong (λ f → P.cong (map f) (concatMap-pure as)) fs ⟨ concatMap (λ f → map f (concatMap pure as)) fs ≡⟨ concatMap-cong (λ f → map-concatMap f pure as) fs ⟩ concatMap (λ f → concatMap (λ x → pure (f x)) as) fs @@ -224,7 +224,7 @@ module Applicative where right-distributive fs₁ fs₂ xs = begin (fs₁ ∣ fs₂) ⊛ xs ≡⟨ unfold-⊛ (fs₁ ∣ fs₂) xs ⟩ (fs₁ ∣ fs₂ >>= pam xs) ≡⟨ MonadProperties.right-distributive fs₁ fs₂ (pam xs) ⟩ - (fs₁ >>= pam xs) ∣ (fs₂ >>= pam xs) ≡˘⟨ P.cong₂ _∣_ (unfold-⊛ fs₁ xs) (unfold-⊛ fs₂ xs) ⟩ + (fs₁ >>= pam xs) ∣ (fs₂ >>= pam xs) ≡⟨ P.cong₂ _∣_ (unfold-⊛ fs₁ xs) (unfold-⊛ fs₂ xs) ⟨ (fs₁ ⊛ xs ∣ fs₂ ⊛ xs) ∎ -- _⊛_ does not distribute over _∣_ from the left. @@ -251,7 +251,7 @@ module Applicative where (xs : List A) (f : A → B) (fs : B → List C) → (pam xs f >>= fs) ≡ (xs >>= λ x → fs (f x)) pam-lemma xs f fs = begin - (pam xs f >>= fs) ≡˘⟨ MP.associative xs (pure ∘ f) fs ⟩ + (pam xs f >>= fs) ≡⟨ MP.associative xs (pure ∘ f) fs ⟨ (xs >>= λ x → pure (f x) >>= fs) ≡⟨ MP.cong (refl {x = xs}) (λ x → MP.left-identity (f x) fs) ⟩ (xs >>= λ x → fs (f x)) ∎ @@ -272,7 +272,7 @@ module Applicative where (pam fs _∘′_ >>= pam gs >>= pam xs) ≡⟨ MP.cong (pam-lemma fs _∘′_ (pam gs)) (λ _ → refl) ⟩ ((fs >>= λ f → pam gs (f ∘′_)) >>= pam xs) - ≡˘⟨ MP.associative fs (λ f → pam gs (_∘′_ f)) (pam xs) ⟩ + ≡⟨ MP.associative fs (λ f → pam gs (_∘′_ f)) (pam xs) ⟨ (fs >>= λ f → pam gs (f ∘′_) >>= pam xs) ≡⟨ MP.cong (refl {x = fs}) (λ f → pam-lemma gs (f ∘′_) (pam xs)) ⟩ (fs >>= λ f → gs >>= λ g → pam xs (f ∘′ g)) @@ -282,9 +282,9 @@ module Applicative where (fs >>= λ f → gs >>= λ g → pam (pam xs g) f) ≡⟨ MP.cong (refl {x = fs}) (λ f → MP.associative gs (pam xs) (pure ∘ f)) ⟩ (fs >>= pam (gs >>= pam xs)) - ≡˘⟨ unfold-⊛ fs (gs >>= pam xs) ⟩ + ≡⟨ unfold-⊛ fs (gs >>= pam xs) ⟨ fs ⊛ (gs >>= pam xs) - ≡˘⟨ P.cong (fs ⊛_) (unfold-⊛ gs xs) ⟩ + ≡⟨ P.cong (fs ⊛_) (unfold-⊛ gs xs) ⟨ fs ⊛ (gs ⊛ xs) ∎ @@ -304,5 +304,5 @@ module Applicative where MP.left-identity x (pure ∘ f)) ⟩ (fs >>= λ f → pure (f x)) ≡⟨⟩ (pam fs (_$′ x)) ≡⟨ P.sym $ MP.left-identity (_$′ x) (pam fs) ⟩ - (pure (_$′ x) >>= pam fs) ≡˘⟨ unfold-⊛ (pure (_$′ x)) fs ⟩ + (pure (_$′ x) >>= pam fs) ≡⟨ unfold-⊛ (pure (_$′ x)) fs ⟨ pure (_$′ x) ⊛ fs ∎ diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda index 78fc03f865..5bd47dc763 100644 --- a/src/Data/List/Properties.agda +++ b/src/Data/List/Properties.agda @@ -280,7 +280,7 @@ module _ (f : A → B → C) where cartesianProductWith-distribʳ-++ (x ∷ xs) ys zs = begin prod (x ∷ xs ++ ys) zs ≡⟨⟩ map (f x) zs ++ prod (xs ++ ys) zs ≡⟨ cong (map (f x) zs ++_) (cartesianProductWith-distribʳ-++ xs ys zs) ⟩ - map (f x) zs ++ prod xs zs ++ prod ys zs ≡˘⟨ ++-assoc (map (f x) zs) (prod xs zs) (prod ys zs) ⟩ + map (f x) zs ++ prod xs zs ++ prod ys zs ≡⟨ ++-assoc (map (f x) zs) (prod xs zs) (prod ys zs) ⟨ (map (f x) zs ++ prod xs zs) ++ prod ys zs ≡⟨⟩ prod (x ∷ xs) zs ++ prod ys zs ∎ @@ -590,7 +590,7 @@ map-concatMap f g xs = begin map f (concatMap g xs) ≡⟨⟩ map f (concat (map g xs)) - ≡˘⟨ concat-map (map g xs) ⟩ + ≡⟨ concat-map (map g xs) ⟨ concat (map (map f) (map g xs)) ≡⟨ cong concat {x = map (map f) (map g xs)} diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index f842ae8087..31ab6941ac 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -96,7 +96,7 @@ module ⊆-Reasoning {A : Set a} where private module Base = PreorderReasoning (⊆-preorder A) open Base public - hiding (step-≈; step-≈˘; step-∼; step-≲) + hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨; step-∼; step-≲) renaming (≲-go to ⊆-go) open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → begin x) public @@ -199,7 +199,7 @@ module _ {k} {A B : Set a} {fs gs : List (A → B)} {xs ys} where x ∈ (fs >>= λ f → xs >>= λ x → pure (f x)) ∼⟨ >>=-cong fs≈gs (λ f → >>=-cong xs≈ys λ x → K-refl) ⟩ x ∈ (gs >>= λ g → ys >>= λ y → pure (g y)) - ≡˘⟨ P.cong (x ∈_) (Applicative.unfold-⊛ gs ys) ⟩ + ≡⟨ P.cong (x ∈_) (Applicative.unfold-⊛ gs ys) ⟨ x ∈ (gs ⊛ ys) ∎ where open Related.EquationalReasoning @@ -284,7 +284,7 @@ empty-unique {xs = _ ∷ _} ∷∼[] with ⇒→ ∷∼[] (here refl) (xs₂ >>= pure ∘ f)) ≈⟨ >>=-left-distributive fs ⟩ (fs >>= λ f → xs₁ >>= pure ∘ f) ++ - (fs >>= λ f → xs₂ >>= pure ∘ f) ≡˘⟨ P.cong₂ _++_ (Applicative.unfold-⊛ fs xs₁) (Applicative.unfold-⊛ fs xs₂) ⟩ + (fs >>= λ f → xs₂ >>= pure ∘ f) ≡⟨ P.cong₂ _++_ (Applicative.unfold-⊛ fs xs₁) (Applicative.unfold-⊛ fs xs₂) ⟨ (fs ⊛ xs₁) ++ (fs ⊛ xs₂) ∎ where open EqR ([ bag ]-Equality B) @@ -592,7 +592,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ... | zs₁ , zs₂ , p rewrite p = begin x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩ x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩ - x ∷ (zs₁ ++ zs₂) ↭˘⟨ shift x zs₁ zs₂ ⟩ + x ∷ (zs₁ ++ zs₂) ↭⟨ shift x zs₁ zs₂ ⟨ zs₁ ++ x ∷ zs₂ ∎ where open CommutativeMonoid (commutativeMonoid bag A) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index 28cf25d9ed..c41793d28a 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -76,9 +76,11 @@ module PermutationReasoning where private module Base = EqReasoning ↭-setoid - open Base public hiding (step-≈; step-≈˘) + open Base public + hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨) + renaming (≈-go to ↭-go) - open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go ↭-sym public + open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ ↭-go ↭-sym public -- Some extra combinators that allow us to skip certain elements diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 1b33a771a0..781de1cc36 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -243,7 +243,7 @@ drop-mid {A = A} {x} ws xs p = drop-mid′ p ws xs refl refl ++-comm [] ys = ↭-sym (++-identityʳ ys) ++-comm (x ∷ xs) ys = begin x ∷ xs ++ ys <⟨ ++-comm xs ys ⟩ - x ∷ ys ++ xs ↭˘⟨ shift x ys xs ⟩ + x ∷ ys ++ xs ↭⟨ shift x ys xs ⟨ ys ++ (x ∷ xs) ∎ ++-isMagma : IsMagma {A = List A} _↭_ _++_ @@ -296,7 +296,7 @@ module _ {a} {A : Set a} where shifts : ∀ xs ys {zs : List A} → xs ++ ys ++ zs ↭ ys ++ xs ++ zs shifts xs ys {zs} = begin - xs ++ ys ++ zs ↭˘⟨ ++-assoc xs ys zs ⟩ + xs ++ ys ++ zs ↭⟨ ++-assoc xs ys zs ⟨ (xs ++ ys) ++ zs ↭⟨ ++⁺ʳ zs (++-comm xs ys) ⟩ (ys ++ xs) ++ zs ↭⟨ ++-assoc ys xs zs ⟩ ys ++ xs ++ zs ∎ @@ -330,7 +330,7 @@ drop-∷ = drop-mid [] [] ↭-reverse [] = ↭-refl ↭-reverse (x ∷ xs) = begin reverse (x ∷ xs) ≡⟨ Lₚ.unfold-reverse x xs ⟩ - reverse xs ∷ʳ x ↭˘⟨ ∷↭∷ʳ x (reverse xs) ⟩ + reverse xs ∷ʳ x ↭⟨ ∷↭∷ʳ x (reverse xs) ⟨ x ∷ reverse xs ↭⟨ prep x (↭-reverse xs) ⟩ x ∷ xs ∎ where open PermutationReasoning @@ -349,7 +349,7 @@ module _ {ℓ} {R : Rel A ℓ} (R? : Decidable R) where ... | true | rec | _ = prep x rec ... | false | _ | rec = begin y ∷ merge R? (x ∷ xs) ys <⟨ rec ⟩ - y ∷ x ∷ xs ++ ys ↭˘⟨ shift y (x ∷ xs) ys ⟩ - (x ∷ xs) ++ y ∷ ys ≡˘⟨ Lₚ.++-assoc [ x ] xs (y ∷ ys) ⟩ + y ∷ x ∷ xs ++ ys ↭⟨ shift y (x ∷ xs) ys ⟨ + (x ∷ xs) ++ y ∷ ys ≡⟨ Lₚ.++-assoc [ x ] xs (y ∷ ys) ⟨ x ∷ xs ++ y ∷ ys ∎ where open PermutationReasoning diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index 0d375226db..edcf62b6da 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -90,10 +90,12 @@ module PermutationReasoning where private module Base = SetoidReasoning ↭-setoid - open Base public hiding (step-≈; step-≈˘) + open Base public + hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨) + renaming (≈-go to ↭-go) - open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go ↭-sym public - open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ (Base.∼-go ∘′ refl) ≋-sym public + open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ ↭-go ↭-sym public + open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ (↭-go ∘′ refl) ≋-sym public -- Some extra combinators that allow us to skip certain elements diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index e09251caf5..3af753e826 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -200,7 +200,7 @@ shift {v} {w} v≈w (x ∷ xs) ys = begin ++-comm [] ys = ↭-sym (++-identityʳ ys) ++-comm (x ∷ xs) ys = begin x ∷ xs ++ ys <⟨ ++-comm xs ys ⟩ - x ∷ ys ++ xs ↭˘⟨ ↭-shift ys xs ⟩ + x ∷ ys ++ xs ↭⟨ ↭-shift ys xs ⟨ ys ++ (x ∷ xs) ∎ -- Structures @@ -262,7 +262,7 @@ inject v ws↭ys xs↭zs = trans (++⁺ˡ _ (↭-prep _ xs↭zs)) (++⁺ʳ _ ws shifts : ∀ xs ys {zs : List A} → xs ++ ys ++ zs ↭ ys ++ xs ++ zs shifts xs ys {zs} = begin - xs ++ ys ++ zs ↭˘⟨ ++-assoc xs ys zs ⟩ + xs ++ ys ++ zs ↭⟨ ++-assoc xs ys zs ⟨ (xs ++ ys) ++ zs ↭⟨ ++⁺ʳ zs (++-comm xs ys) ⟩ (ys ++ xs) ++ zs ↭⟨ ++-assoc ys xs zs ⟩ ys ++ xs ++ zs ∎ @@ -297,19 +297,19 @@ dropMiddleElement {v} ws xs {ys} {zs} p = helper p ws xs ≋-refl ≋-refl helper {_ ∷ as} {_ ∷ bs} (prep _ as↭bs) [] [] {ys} {zs} (_ ∷ ys≋as) (_ ∷ zs≋bs) = begin ys ≋⟨ ys≋as ⟩ as ↭⟨ as↭bs ⟩ - bs ≋˘⟨ zs≋bs ⟩ + bs ≋⟨ zs≋bs ⟨ zs ∎ helper {_ ∷ as} {_ ∷ bs} (prep a≈b as↭bs) [] (x ∷ xs) {ys} {zs} (≈₁ ∷ ≋₁) (≈₂ ∷ ≋₂) = begin ys ≋⟨ ≋₁ ⟩ as ↭⟨ as↭bs ⟩ - bs ≋˘⟨ ≋₂ ⟩ + bs ≋⟨ ≋₂ ⟨ xs ++ v ∷ zs ↭⟨ shift (lemma ≈₁ a≈b ≈₂) xs zs ⟩ x ∷ xs ++ zs ∎ helper {_ ∷ as} {_ ∷ bs} (prep v≈w p) (w ∷ ws) [] {ys} {zs} (≈₁ ∷ ≋₁) (≈₂ ∷ ≋₂) = begin w ∷ ws ++ ys ↭⟨ ↭-sym (shift (lemma ≈₂ (≈-sym v≈w) ≈₁) ws ys) ⟩ ws ++ v ∷ ys ≋⟨ ≋₁ ⟩ as ↭⟨ p ⟩ - bs ≋˘⟨ ≋₂ ⟩ + bs ≋⟨ ≋₂ ⟨ zs ∎ helper {_ ∷ as} {_ ∷ bs} (prep w≈x p) (w ∷ ws) (x ∷ xs) {ys} {zs} (≈₁ ∷ ≋₁) (≈₂ ∷ ≋₂) = begin w ∷ ws ++ ys ↭⟨ prep (lemma ≈₁ w≈x ≈₂) (helper p ws xs ≋₁ ≋₂) ⟩ @@ -317,12 +317,12 @@ dropMiddleElement {v} ws xs {ys} {zs} p = helper p ws xs ≋-refl ≋-refl helper {_ ∷ a ∷ as} {_ ∷ b ∷ bs} (swap v≈x y≈v p) [] [] {ys} {zs} (≈₁ ∷ ≋₁) (≈₂ ∷ ≋₂) = begin ys ≋⟨ ≋₁ ⟩ a ∷ as ↭⟨ prep (≈-trans (≈-trans (≈-trans y≈v (≈-sym ≈₂)) ≈₁) v≈x) p ⟩ - b ∷ bs ≋˘⟨ ≋₂ ⟩ + b ∷ bs ≋⟨ ≋₂ ⟨ zs ∎ helper {_ ∷ a ∷ as} {_ ∷ b ∷ bs} (swap v≈w y≈w p) [] (x ∷ []) {ys} {zs} (≈₁ ∷ ≋₁) (≈₂ ∷ ≋₂) = begin ys ≋⟨ ≋₁ ⟩ a ∷ as ↭⟨ prep y≈w p ⟩ - _ ∷ bs ≋˘⟨ ≈₂ ∷ tail ≋₂ ⟩ + _ ∷ bs ≋⟨ ≈₂ ∷ tail ≋₂ ⟨ x ∷ zs ∎ helper {_ ∷ a ∷ as} {_ ∷ b ∷ bs} (swap v≈w y≈x p) [] (x ∷ w ∷ xs) {ys} {zs} (≈₁ ∷ ≋₁) (≈₂ ∷ ≋₂) = begin ys ≋⟨ ≋₁ ⟩ @@ -455,8 +455,8 @@ module _ {ℓ} {R : Rel A ℓ} (R? : B.Decidable R) where ... | true | rec | _ = ↭-prep x rec ... | false | _ | rec = begin y ∷ merge R? (x ∷ xs) ys <⟨ rec ⟩ - y ∷ x ∷ xs ++ ys ↭˘⟨ ↭-shift (x ∷ xs) ys ⟩ - (x ∷ xs) ++ y ∷ ys ≡˘⟨ Lₚ.++-assoc [ x ] xs (y ∷ ys) ⟩ + y ∷ x ∷ xs ++ ys ↭⟨ ↭-shift (x ∷ xs) ys ⟨ + (x ∷ xs) ++ y ∷ ys ≡⟨ Lₚ.++-assoc [ x ] xs (y ∷ ys) ⟨ x ∷ xs ++ y ∷ ys ∎ where open PermutationReasoning @@ -497,7 +497,7 @@ module _ {_∙_ : Op₂ A} {ε : A} (isCmonoid : IsCommutativeMonoid _≈_ _∙_ foldr-commMonoid (prep x≈y xs↭ys) = ∙-cong x≈y (foldr-commMonoid xs↭ys) foldr-commMonoid (swap {xs} {ys} {x} {y} {x′} {y′} x≈x′ y≈y′ xs↭ys) = S.begin x ∙ (y ∙ foldr _∙_ ε xs) S.≈⟨ ∙-congˡ (∙-congˡ (foldr-commMonoid xs↭ys)) ⟩ - x ∙ (y ∙ foldr _∙_ ε ys) S.≈˘⟨ assoc x y (foldr _∙_ ε ys) ⟩ + x ∙ (y ∙ foldr _∙_ ε ys) S.≈⟨ assoc x y (foldr _∙_ ε ys) ⟨ (x ∙ y) ∙ foldr _∙_ ε ys S.≈⟨ ∙-congʳ (comm x y) ⟩ (y ∙ x) ∙ foldr _∙_ ε ys S.≈⟨ ∙-congʳ (∙-cong y≈y′ x≈x′) ⟩ (y′ ∙ x′) ∙ foldr _∙_ ε ys S.≈⟨ assoc y′ x′ (foldr _∙_ ε ys) ⟩ diff --git a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda index bb20fd417f..0c344758b7 100644 --- a/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda @@ -118,7 +118,7 @@ module ⊆-Reasoning (S : Setoid a ℓ) where private module Base = PreorderReasoning (⊆-preorder S) open Base public - hiding (step-≲; step-≈; step-≈˘; step-∼) + hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨; step-≲; step-∼) renaming (≲-go to ⊆-go; ≈-go to ≋-go) open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x → Base.begin x) public diff --git a/src/Data/List/Relation/Unary/Any/Properties.agda b/src/Data/List/Relation/Unary/Any/Properties.agda index b88bdfbe95..e356e9eca1 100644 --- a/src/Data/List/Relation/Unary/Any/Properties.agda +++ b/src/Data/List/Relation/Unary/Any/Properties.agda @@ -686,7 +686,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where Any (λ f → Any (P ∘ f) xs) fs ↔⟨ Any-cong (λ _ → Any-cong (λ _ → pure↔) (_ ∎)) (_ ∎) ⟩ Any (λ f → Any (Any P ∘ pure ∘ f) xs) fs ↔⟨ Any-cong (λ _ → >>=↔ ) (_ ∎) ⟩ Any (λ f → Any P (xs >>= pure ∘ f)) fs ↔⟨ >>=↔ ⟩ - Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡˘⟨ P.cong (Any P) (Listₑ.Applicative.unfold-⊛ fs xs) ⟩ + Any P (fs >>= λ f → xs >>= λ x → pure (f x)) ≡⟨ P.cong (Any P) (Listₑ.Applicative.unfold-⊛ fs xs) ⟨ Any P (fs ⊛ xs) ∎ where open Related.EquationalReasoning @@ -706,7 +706,7 @@ module _ {A B : Set ℓ} {P : B → Set p} {f : A → List B} where Any (λ x → Any (λ y → P (x , y)) ys) xs ↔⟨ pure↔ ⟩ Any (λ _,_ → Any (λ x → Any (λ y → P (x , y)) ys) xs) (pure _,_) ↔⟨ ⊛↔ ⟩ Any (λ x, → Any (P ∘ x,) ys) (pure _,_ ⊛ xs) ↔⟨ ⊛↔ ⟩ - Any P (pure _,_ ⊛ xs ⊛ ys) ≡˘⟨ P.cong (Any P ∘′ (_⊛ ys)) (Listₑ.Applicative.unfold-<$> _,_ xs) ⟩ + Any P (pure _,_ ⊛ xs ⊛ ys) ≡⟨ P.cong (Any P ∘′ (_⊛ ys)) (Listₑ.Applicative.unfold-<$> _,_ xs) ⟨ Any P (xs ⊗ ys) ∎ where open Related.EquationalReasoning diff --git a/src/Data/Nat/Binary/Properties.agda b/src/Data/Nat/Binary/Properties.agda index 9e8976842e..f83bbfc12c 100644 --- a/src/Data/Nat/Binary/Properties.agda +++ b/src/Data/Nat/Binary/Properties.agda @@ -156,9 +156,9 @@ fromℕ≡fromℕ' n = fromℕ-helper≡fromℕ' n n ℕₚ.≤-refl tail-homo ℕ.zero = refl tail-homo (ℕ.suc ℕ.zero) = refl tail-homo (ℕ.suc (ℕ.suc n)) = begin - tail (suc (fromℕ' (ℕ.suc (ℕ.suc n)))) ≡˘⟨ tail-suc (fromℕ' n) ⟩ + tail (suc (fromℕ' (ℕ.suc (ℕ.suc n)))) ≡⟨ tail-suc (fromℕ' n) ⟨ suc (tail (suc (fromℕ' n))) ≡⟨ cong suc (tail-homo n) ⟩ - fromℕ' (ℕ.suc (n / 2)) ≡˘⟨ cong fromℕ' (+-distrib-/-∣ˡ {2} n ∣-refl) ⟩ + fromℕ' (ℕ.suc (n / 2)) ≡⟨ cong fromℕ' (+-distrib-/-∣ˡ {2} n ∣-refl) ⟨ fromℕ' (ℕ.suc (ℕ.suc n) / 2) ∎ fromℕ-helper≡fromℕ' : ∀ n w → n ℕ.≤ w → fromℕ.helper n n w ≡ fromℕ' n @@ -167,7 +167,7 @@ fromℕ≡fromℕ' n = fromℕ-helper≡fromℕ' n n ℕₚ.≤-refl split-injective (begin split (fromℕ.helper n (ℕ.suc n) (ℕ.suc w)) ≡⟨ split-if _ _ ⟩ just (n % 2 ℕ.≡ᵇ 0) , fromℕ.helper n (n / 2) w ≡⟨ cong (_ ,_) rec-n/2 ⟩ - just (n % 2 ℕ.≡ᵇ 0) , fromℕ' (n / 2) ≡˘⟨ cong₂ _,_ (head-homo n) (tail-homo n) ⟩ + just (n % 2 ℕ.≡ᵇ 0) , fromℕ' (n / 2) ≡⟨ cong₂ _,_ (head-homo n) (tail-homo n) ⟨ head (fromℕ' (ℕ.suc n)) , tail (fromℕ' (ℕ.suc n)) ≡⟨⟩ split (fromℕ' (ℕ.suc n)) ∎) where rec-n/2 = fromℕ-helper≡fromℕ' (n / 2) w (ℕₚ.≤-trans (m/n≤m n 2) n≤w) @@ -613,7 +613,7 @@ module ≤-Reasoning where <-≤-trans ≤-<-trans public - hiding (step-≈; step-≈˘) + hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨) ------------------------------------------------------------------------ -- Properties of _<ℕ_ diff --git a/src/Data/Nat/Combinatorics.agda b/src/Data/Nat/Combinatorics.agda index fc6ef422b4..d07d22db47 100644 --- a/src/Data/Nat/Combinatorics.agda +++ b/src/Data/Nat/Combinatorics.agda @@ -66,9 +66,9 @@ open Specification public nCk≡nC[n∸k] : k ≤ n → n C k ≡ n C (n ∸ k) nCk≡nC[n∸k] {k} {n} k≤n = begin-equality n C k ≡⟨ nCk≡n!/k![n-k]! k≤n ⟩ - n ! / (k ! * (n ∸ k) !) ≡˘⟨ /-congʳ (*-comm ((n ∸ k) !) (k !)) ⟩ - n ! / ((n ∸ k) ! * k !) ≡˘⟨ /-congʳ (cong ((n ∸ k) ! *_) (cong _! (m∸[m∸n]≡n k≤n))) ⟩ - n ! / ((n ∸ k) ! * (n ∸ (n ∸ k)) !) ≡˘⟨ nCk≡n!/k![n-k]! (m≤n⇒n∸m≤n k≤n) ⟩ + n ! / (k ! * (n ∸ k) !) ≡⟨ /-congʳ (*-comm ((n ∸ k) !) (k !)) ⟨ + n ! / ((n ∸ k) ! * k !) ≡⟨ /-congʳ (cong ((n ∸ k) ! *_) (cong _! (m∸[m∸n]≡n k≤n))) ⟨ + n ! / ((n ∸ k) ! * (n ∸ (n ∸ k)) !) ≡⟨ nCk≡n!/k![n-k]! (m≤n⇒n∸m≤n k≤n) ⟨ n C (n ∸ k) ∎ where instance _ = (n ∸ k) !* k !≢0 @@ -78,9 +78,9 @@ nCk≡nC[n∸k] {k} {n} k≤n = begin-equality nCk≡nPk/k! : k ≤ n → n C k ≡ ((n P k) / k !) {{k !≢0}} nCk≡nPk/k! {k} {n} k≤n = begin-equality n C k ≡⟨ nCk≡n!/k![n-k]! k≤n ⟩ - n ! / (k ! * (n ∸ k) !) ≡˘⟨ /-congʳ (*-comm ((n ∸ k)!) (k !)) ⟩ - n ! / ((n ∸ k) ! * k !) ≡˘⟨ m/n/o≡m/[n*o] (n !) ((n ∸ k) !) (k !) ⟩ - (n ! / (n ∸ k) !) / k ! ≡˘⟨ /-congˡ (nPk≡n!/[n∸k]! k≤n) ⟩ + n ! / (k ! * (n ∸ k) !) ≡⟨ /-congʳ (*-comm ((n ∸ k)!) (k !)) ⟨ + n ! / ((n ∸ k) ! * k !) ≡⟨ m/n/o≡m/[n*o] (n !) ((n ∸ k) !) (k !) ⟨ + (n ! / (n ∸ k) !) / k ! ≡⟨ /-congˡ (nPk≡n!/[n∸k]! k≤n) ⟨ (n P k) / k ! ∎ where instance _ = k !≢0 @@ -128,7 +128,7 @@ module _ {n k} (kn⇒nCk≡0 k>n) ⟩ 0 + n C (suc k) ≡⟨⟩ n C (suc k) ≡⟨ k>n⇒nCk≡0 (mn) ⟩ - 0 ≡˘⟨ k>n⇒nCk≡0 (sn) ⟩ + 0 ≡⟨ k>n⇒nCk≡0 (sn) ⟨ suc n C suc k ∎ {- case k≡n, in which case 1+k≡1+n>n -} ... | tri≈ _ k≡n _ rewrite k≡n = begin-equality n C n + n C (suc n) ≡⟨ cong (n C n +_) (k>n⇒nCk≡0 (n<1+n n)) ⟩ n C n + 0 ≡⟨ +-identityʳ (n C n) ⟩ n C n ≡⟨ nCn≡1 n ⟩ - 1 ≡˘⟨ nCn≡1 (suc n) ⟩ + 1 ≡⟨ nCn≡1 (suc n) ⟨ suc n C suc n ∎ {- case k0 (≢-nonZero⁻¹ o)) (≮⇒≥ n≮o) @@ -333,19 +333,19 @@ m q *-cancelˡ-<-nonPos {p} {q} r rp p≮q p≡r p>q = begin f (p ⊓ q) ≡⟨ cong f (p≥q⇒p⊓q≡q (<⇒≤ p>q)) ⟩ - f q ≡˘⟨ p≥q⇒p⊓q≡q (<⇒≤ (f-mono-< p>q)) ⟩ + f q ≡⟨ p≥q⇒p⊓q≡q (<⇒≤ (f-mono-< p>q)) ⟨ f p ⊓ f q ∎ where open ≡-Reasoning @@ -1507,17 +1508,17 @@ mono-<-distrib-⊔ : ∀ {f} → f Preserves _<_ ⟶ _<_ → mono-<-distrib-⊔ {f} f-mono-< p q with <-cmp p q ... | tri< p p≮q p≡r p>q = begin f (p ⊔ q) ≡⟨ cong f (p≥q⇒p⊔q≡p (<⇒≤ p>q)) ⟩ - f p ≡˘⟨ p≥q⇒p⊔q≡p (<⇒≤ (f-mono-< p>q)) ⟩ + f p ≡⟨ p≥q⇒p⊔q≡p (<⇒≤ (f-mono-< p>q)) ⟨ f p ⊔ f q ∎ where open ≡-Reasoning @@ -1605,7 +1606,7 @@ toℚᵘ-homo-∣-∣ (mkℚ -[1+ _ ] _ _) = *≡* refl 0≤∣p∣ p@record{} = *≤* (begin (↥ 0ℚ) ℤ.* (↧ ∣ p ∣) ≡⟨ ℤ.*-zeroˡ (↧ ∣ p ∣) ⟩ 0ℤ ≤⟨ ℤ.+≤+ ℕ.z≤n ⟩ - ↥ ∣ p ∣ ≡˘⟨ ℤ.*-identityʳ (↥ ∣ p ∣) ⟩ + ↥ ∣ p ∣ ≡⟨ ℤ.*-identityʳ (↥ ∣ p ∣) ⟨ ↥ ∣ p ∣ ℤ.* 1ℤ ∎) where open ℤ.≤-Reasoning @@ -1633,8 +1634,8 @@ toℚᵘ-homo-∣-∣ (mkℚ -[1+ _ ] _ _) = *≡* refl toℚᵘ ∣ p + q ∣ ≃⟨ toℚᵘ-homo-∣-∣ (p + q) ⟩ ℚᵘ.∣ toℚᵘ (p + q) ∣ ≃⟨ ℚᵘ.∣-∣-cong (toℚᵘ-homo-+ p q) ⟩ ℚᵘ.∣ toℚᵘ p ℚᵘ.+ toℚᵘ q ∣ ≤⟨ ℚᵘ.∣p+q∣≤∣p∣+∣q∣ (toℚᵘ p) (toℚᵘ q) ⟩ - ℚᵘ.∣ toℚᵘ p ∣ ℚᵘ.+ ℚᵘ.∣ toℚᵘ q ∣ ≃˘⟨ ℚᵘ.+-cong (toℚᵘ-homo-∣-∣ p) (toℚᵘ-homo-∣-∣ q) ⟩ - toℚᵘ ∣ p ∣ ℚᵘ.+ toℚᵘ ∣ q ∣ ≃˘⟨ toℚᵘ-homo-+ ∣ p ∣ ∣ q ∣ ⟩ + ℚᵘ.∣ toℚᵘ p ∣ ℚᵘ.+ ℚᵘ.∣ toℚᵘ q ∣ ≃⟨ ℚᵘ.+-cong (toℚᵘ-homo-∣-∣ p) (toℚᵘ-homo-∣-∣ q) ⟨ + toℚᵘ ∣ p ∣ ℚᵘ.+ toℚᵘ ∣ q ∣ ≃⟨ toℚᵘ-homo-+ ∣ p ∣ ∣ q ∣ ⟨ toℚᵘ (∣ p ∣ + ∣ q ∣) ∎) where open ℚᵘ.≤-Reasoning @@ -1650,8 +1651,8 @@ toℚᵘ-homo-∣-∣ (mkℚ -[1+ _ ] _ _) = *≡* refl toℚᵘ ∣ p * q ∣ ≃⟨ toℚᵘ-homo-∣-∣ (p * q) ⟩ ℚᵘ.∣ toℚᵘ (p * q) ∣ ≃⟨ ℚᵘ.∣-∣-cong (toℚᵘ-homo-* p q) ⟩ ℚᵘ.∣ toℚᵘ p ℚᵘ.* toℚᵘ q ∣ ≃⟨ ℚᵘ.∣p*q∣≃∣p∣*∣q∣ (toℚᵘ p) (toℚᵘ q) ⟩ - ℚᵘ.∣ toℚᵘ p ∣ ℚᵘ.* ℚᵘ.∣ toℚᵘ q ∣ ≃˘⟨ ℚᵘ.*-cong (toℚᵘ-homo-∣-∣ p) (toℚᵘ-homo-∣-∣ q) ⟩ - toℚᵘ ∣ p ∣ ℚᵘ.* toℚᵘ ∣ q ∣ ≃˘⟨ toℚᵘ-homo-* ∣ p ∣ ∣ q ∣ ⟩ + ℚᵘ.∣ toℚᵘ p ∣ ℚᵘ.* ℚᵘ.∣ toℚᵘ q ∣ ≃⟨ ℚᵘ.*-cong (toℚᵘ-homo-∣-∣ p) (toℚᵘ-homo-∣-∣ q) ⟨ + toℚᵘ ∣ p ∣ ℚᵘ.* toℚᵘ ∣ q ∣ ≃⟨ toℚᵘ-homo-* ∣ p ∣ ∣ q ∣ ⟨ toℚᵘ (∣ p ∣ * ∣ q ∣) ∎) where open ℚᵘ.≤-Reasoning diff --git a/src/Data/Rational/Unnormalised/Properties.agda b/src/Data/Rational/Unnormalised/Properties.agda index 6348a6ef46..dce11c3f21 100644 --- a/src/Data/Rational/Unnormalised/Properties.agda +++ b/src/Data/Rational/Unnormalised/Properties.agda @@ -160,7 +160,7 @@ module ≃-Reasoning = SetoidReasoning ≃-setoid p≃0⇒↥p≡0 : ∀ p → p ≃ 0ℚᵘ → ↥ p ≡ 0ℤ p≃0⇒↥p≡0 p (*≡* eq) = begin - ↥ p ≡˘⟨ ℤ.*-identityʳ (↥ p) ⟩ + ↥ p ≡⟨ ℤ.*-identityʳ (↥ p) ⟨ ↥ p ℤ.* 1ℤ ≡⟨ eq ⟩ 0ℤ ∎ where open ≡-Reasoning @@ -180,14 +180,14 @@ neg-involutive p rewrite neg-involutive-≡ p = ≃-refl -‿cong : Congruent₁ _≃_ (-_) -‿cong {p@record{}} {q@record{}} (*≡* p≡q) = *≡* (begin - ↥(- p) ℤ.* ↧ q ≡˘⟨ ℤ.*-identityˡ (ℤ.- (↥ p) ℤ.* ↧ q) ⟩ - 1ℤ ℤ.* (↥(- p) ℤ.* ↧ q) ≡˘⟨ ℤ.*-assoc 1ℤ (↥ (- p)) (↧ q) ⟩ - (1ℤ ℤ.* ℤ.-(↥ p)) ℤ.* ↧ q ≡˘⟨ cong (ℤ._* ↧ q) (ℤ.neg-distribʳ-* 1ℤ (↥ p)) ⟩ + ↥(- p) ℤ.* ↧ q ≡⟨ ℤ.*-identityˡ (ℤ.- (↥ p) ℤ.* ↧ q) ⟨ + 1ℤ ℤ.* (↥(- p) ℤ.* ↧ q) ≡⟨ ℤ.*-assoc 1ℤ (↥ (- p)) (↧ q) ⟨ + (1ℤ ℤ.* ℤ.-(↥ p)) ℤ.* ↧ q ≡⟨ cong (ℤ._* ↧ q) (ℤ.neg-distribʳ-* 1ℤ (↥ p)) ⟨ ℤ.-(1ℤ ℤ.* ↥ p) ℤ.* ↧ q ≡⟨ cong (ℤ._* ↧ q) (ℤ.neg-distribˡ-* 1ℤ (↥ p)) ⟩ (-1ℤ ℤ.* ↥ p) ℤ.* ↧ q ≡⟨ ℤ.*-assoc ℤ.-1ℤ (↥ p) (↧ q) ⟩ -1ℤ ℤ.* (↥ p ℤ.* ↧ q) ≡⟨ cong (ℤ.-1ℤ ℤ.*_) p≡q ⟩ - -1ℤ ℤ.* (↥ q ℤ.* ↧ p) ≡˘⟨ ℤ.*-assoc ℤ.-1ℤ (↥ q) (↧ p) ⟩ - (-1ℤ ℤ.* ↥ q) ℤ.* ↧ p ≡˘⟨ cong (ℤ._* ↧ p) (ℤ.neg-distribˡ-* 1ℤ (↥ q)) ⟩ + -1ℤ ℤ.* (↥ q ℤ.* ↧ p) ≡⟨ ℤ.*-assoc ℤ.-1ℤ (↥ q) (↧ p) ⟨ + (-1ℤ ℤ.* ↥ q) ℤ.* ↧ p ≡⟨ cong (ℤ._* ↧ p) (ℤ.neg-distribˡ-* 1ℤ (↥ q)) ⟨ ℤ.-(1ℤ ℤ.* ↥ q) ℤ.* ↧ p ≡⟨ cong (ℤ._* ↧ p) (ℤ.neg-distribʳ-* 1ℤ (↥ q)) ⟩ (1ℤ ℤ.* ↥(- q)) ℤ.* ↧ p ≡⟨ ℤ.*-assoc 1ℤ (ℤ.- (↥ q)) (↧ p) ⟩ 1ℤ ℤ.* (↥(- q) ℤ.* ↧ p) ≡⟨ ℤ.*-identityˡ (↥ (- q) ℤ.* ↧ p) ⟩ @@ -196,7 +196,7 @@ neg-involutive p rewrite neg-involutive-≡ p = ≃-refl neg-mono-< : -_ Preserves _<_ ⟶ _>_ neg-mono-< {p@record{}} {q@record{}} (*<* p?_ = flip __ *-monoˡ-<-neg r {p} {q} p0}} p0 = positive (neg-mono-< (negative⁻¹ r)) @@ -1328,7 +1329,7 @@ private *-cancelˡ-<-nonPos : ∀ r .{{_ : NonPositive r}} → r * p < r * q → q < p *-cancelˡ-<-nonPos {p} {q} r rp1⇒1/p<1 {p} p>1 = lemma′ p (p>1⇒p≢0 p>1) p>1 1/-antimono-≤-pos : ∀ {p q} .{{_ : Positive p}} .{{_ : Positive q}} → p ≤ q → (1/ q) {{pos⇒nonZero q}} ≤ (1/ p) {{pos⇒nonZero p}} 1/-antimono-≤-pos {p} {q} p≤q = begin - 1/q ≃˘⟨ *-identityˡ 1/q ⟩ - 1ℚᵘ * 1/q ≃˘⟨ *-congʳ (*-inverseˡ p) ⟩ + 1/q ≃⟨ *-identityˡ 1/q ⟨ + 1ℚᵘ * 1/q ≃⟨ *-congʳ (*-inverseˡ p) ⟨ (1/p * p) * 1/q ≤⟨ *-monoˡ-≤-nonNeg 1/q (*-monoʳ-≤-nonNeg 1/p p≤q) ⟩ (1/p * q) * 1/q ≃⟨ *-assoc 1/p q 1/q ⟩ 1/p * (q * 1/q) ≃⟨ *-congˡ {1/p} (*-inverseʳ q) ⟩ @@ -1775,7 +1776,7 @@ pos⊔pos⇒pos p q = positive (⊔-mono-< (positive⁻¹ p) (positive⁻¹ q)) ℤ.- ℤ.- (↥ ∣ p ∣ ℤ.* ↧ q) ≡⟨ cong ℤ.-_ (ℤ.neg-distribˡ-* (↥ ∣ p ∣) (↧ q)) ⟩ ℤ.- (↥ p ℤ.* ↧ q) ≡⟨ cong ℤ.-_ ↥p↧q≡↥q↧p ⟩ ℤ.- (↥ q ℤ.* ↧ p) ≡⟨ cong ℤ.-_ (ℤ.neg-distribˡ-* (↥ ∣ q ∣) (↧ p)) ⟩ - ℤ.- ℤ.- (↥ ∣ q ∣ ℤ.* ↧ p) ≡˘⟨ ℤ.neg-involutive _ ⟩ + ℤ.- ℤ.- (↥ ∣ q ∣ ℤ.* ↧ p) ≡⟨ ℤ.neg-involutive _ ⟨ ↥ ∣ q ∣ ℤ.* ↧ p ∎) where open ≡-Reasoning @@ -1807,7 +1808,7 @@ pos⊔pos⇒pos p q = positive (⊔-mono-< (positive⁻¹ p) (positive⁻¹ q)) ∣p∣≡p⇒0≤p {mkℚᵘ (ℤ.+ n) d-1} ∣p∣≡p = *≤* (begin 0ℤ ℤ.* +[1+ d-1 ] ≡⟨ ℤ.*-zeroˡ (ℤ.+ d-1) ⟩ 0ℤ ≤⟨ ℤ.+≤+ ℕ.z≤n ⟩ - ℤ.+ n ≡˘⟨ ℤ.*-identityʳ (ℤ.+ n) ⟩ + ℤ.+ n ≡⟨ ℤ.*-identityʳ (ℤ.+ n) ⟨ ℤ.+ n ℤ.* 1ℤ ∎) where open ℤ.≤-Reasoning @@ -1826,7 +1827,7 @@ pos⊔pos⇒pos p q = positive (⊔-mono-< (positive⁻¹ p) (positive⁻¹ q)) ↥ ∣ (↥p↧q ℤ.+ ↥q↧p) / ↧p↧q ∣ ℤ.* ℤ.+ ↧p↧q ≡⟨⟩ ↥ (ℤ.+ ℤ.∣ ↥p↧q ℤ.+ ↥q↧p ∣ / ↧p↧q) ℤ.* ℤ.+ ↧p↧q ≡⟨ cong (ℤ._* ℤ.+ ↧p↧q) (↥[n/d]≡n (ℤ.+ ℤ.∣ ↥p↧q ℤ.+ ↥q↧p ∣) ↧p↧q) ⟩ ℤ.+ ℤ.∣ ↥p↧q ℤ.+ ↥q↧p ∣ ℤ.* ℤ.+ ↧p↧q ≤⟨ ℤ.*-monoʳ-≤-nonNeg (ℤ.+ ↧p↧q) (ℤ.+≤+ (ℤ.∣i+j∣≤∣i∣+∣j∣ ↥p↧q ↥q↧p)) ⟩ - (ℤ.+ ℤ.∣ ↥p↧q ∣ ℤ.+ ℤ.+ ℤ.∣ ↥q↧p ∣) ℤ.* ℤ.+ ↧p↧q ≡˘⟨ cong₂ (λ h₁ h₂ → (h₁ ℤ.+ h₂) ℤ.* ℤ.+ ↧p↧q) ∣↥p∣↧q≡∣↥p↧q∣ ∣↥q∣↧p≡∣↥q↧p∣ ⟩ + (ℤ.+ ℤ.∣ ↥p↧q ∣ ℤ.+ ℤ.+ ℤ.∣ ↥q↧p ∣) ℤ.* ℤ.+ ↧p↧q ≡⟨ cong₂ (λ h₁ h₂ → (h₁ ℤ.+ h₂) ℤ.* ℤ.+ ↧p↧q) ∣↥p∣↧q≡∣↥p↧q∣ ∣↥q∣↧p≡∣↥q↧p∣ ⟨ (∣↥p∣↧q ℤ.+ ∣↥q∣↧p) ℤ.* ℤ.+ ↧p↧q ≡⟨⟩ (↥∣p∣↧q ℤ.+ ↥∣q∣↧p) ℤ.* ℤ.+ ↧p↧q ≡⟨ cong (ℤ._* ℤ.+ ↧p↧q) (↥[n/d]≡n (↥∣p∣↧q ℤ.+ ↥∣q∣↧p) ↧p↧q) ⟩ ↥ ((↥∣p∣↧q ℤ.+ ↥∣q∣↧p) / ↧p↧q) ℤ.* ℤ.+ ↧p↧q ≡⟨⟩ @@ -1843,9 +1844,9 @@ pos⊔pos⇒pos p q = positive (⊔-mono-< (positive⁻¹ p) (positive⁻¹ q)) ∣m∣n≡∣mn∣ : ∀ m n → ℤ.+ ℤ.∣ m ∣ ℤ.* ℤ.+ n ≡ ℤ.+ ℤ.∣ m ℤ.* ℤ.+ n ∣ ∣m∣n≡∣mn∣ m n = begin-equality ℤ.+ ℤ.∣ m ∣ ℤ.* ℤ.+ n ≡⟨⟩ - ℤ.+ ℤ.∣ m ∣ ℤ.* ℤ.+ ℤ.∣ ℤ.+ n ∣ ≡˘⟨ ℤ.pos-* ℤ.∣ m ∣ ℤ.∣ ℤ.+ n ∣ ⟩ + ℤ.+ ℤ.∣ m ∣ ℤ.* ℤ.+ ℤ.∣ ℤ.+ n ∣ ≡⟨ ℤ.pos-* ℤ.∣ m ∣ ℤ.∣ ℤ.+ n ∣ ⟨ ℤ.+ (ℤ.∣ m ∣ ℕ.* n) ≡⟨⟩ - ℤ.+ (ℤ.∣ m ∣ ℕ.* ℤ.∣ ℤ.+ n ∣) ≡˘⟨ cong ℤ.+_ (ℤ.∣i*j∣≡∣i∣*∣j∣ m (ℤ.+ n)) ⟩ + ℤ.+ (ℤ.∣ m ∣ ℕ.* ℤ.∣ ℤ.+ n ∣) ≡⟨ cong ℤ.+_ (ℤ.∣i*j∣≡∣i∣*∣j∣ m (ℤ.+ n)) ⟨ ℤ.+ (ℤ.∣ m ℤ.* ℤ.+ n ∣) ∎ ∣↥p∣↧q≡∣↥p↧q∣ : ∣↥p∣↧q ≡ ℤ.+ ℤ.∣ ↥p↧q ∣ ∣↥p∣↧q≡∣↥p↧q∣ = ∣m∣n≡∣mn∣ (↥ p) (↧ₙ q) diff --git a/src/Data/Tree/Binary/Zipper/Properties.agda b/src/Data/Tree/Binary/Zipper/Properties.agda index 107981bcbf..fa2f09c190 100644 --- a/src/Data/Tree/Binary/Zipper/Properties.agda +++ b/src/Data/Tree/Binary/Zipper/Properties.agda @@ -60,9 +60,9 @@ toTree-#nodes (mkZipper c v) = helper c v #ctx = sum (List.map (suc ∘ BT.#nodes ∘ getTree) ctx) #foc = BT.#nodes foc #l = BT.#nodes l in begin - #foc + (1 + (#l + #ctx)) ≡˘⟨ +-assoc #foc 1 (#l + #ctx) ⟩ + #foc + (1 + (#l + #ctx)) ≡⟨ +-assoc #foc 1 (#l + #ctx) ⟨ #foc + 1 + (#l + #ctx) ≡⟨ cong (_+ (#l + #ctx)) (+-comm #foc 1) ⟩ - 1 + #foc + (#l + #ctx) ≡˘⟨ +-assoc (1 + #foc) #l #ctx ⟩ + 1 + #foc + (#l + #ctx) ≡⟨ +-assoc (1 + #foc) #l #ctx ⟨ 1 + #foc + #l + #ctx ≡⟨ cong (_+ #ctx) (+-comm (1 + #foc) #l) ⟩ #nodes (mkZipper ctx (node l m foc)) ≡⟨ helper ctx (node l m foc) ⟩ BT.#nodes (toTree (mkZipper cs foc)) ∎ @@ -70,8 +70,8 @@ toTree-#nodes (mkZipper c v) = helper c v #ctx = sum (List.map (suc ∘ BT.#nodes ∘ getTree) ctx) #foc = BT.#nodes foc #r = BT.#nodes r in begin - #foc + (1 + (#r + #ctx)) ≡˘⟨ cong (#foc +_) (+-assoc 1 #r #ctx) ⟩ - #foc + (1 + #r + #ctx) ≡˘⟨ +-assoc #foc (suc #r) #ctx ⟩ + #foc + (1 + (#r + #ctx)) ≡⟨ cong (#foc +_) (+-assoc 1 #r #ctx) ⟨ + #foc + (1 + #r + #ctx) ≡⟨ +-assoc #foc (suc #r) #ctx ⟨ #nodes (mkZipper ctx (node foc m r)) ≡⟨ helper ctx (node foc m r) ⟩ BT.#nodes (toTree (mkZipper cs foc)) ∎ @@ -86,7 +86,7 @@ toTree-#leaves (mkZipper c v) = helper c v #ctx = sum (List.map (BT.#leaves ∘ getTree) ctx) #foc = BT.#leaves foc #l = BT.#leaves l in begin - #foc + (#l + #ctx) ≡˘⟨ +-assoc #foc #l #ctx ⟩ + #foc + (#l + #ctx) ≡⟨ +-assoc #foc #l #ctx ⟨ #foc + #l + #ctx ≡⟨ cong (_+ #ctx) (+-comm #foc #l) ⟩ #leaves (mkZipper ctx (node l m foc)) ≡⟨ helper ctx (node l m foc) ⟩ BT.#leaves (toTree (mkZipper cs foc)) ∎ @@ -94,7 +94,7 @@ toTree-#leaves (mkZipper c v) = helper c v #ctx = sum (List.map (BT.#leaves ∘ getTree) ctx) #foc = BT.#leaves foc #r = BT.#leaves r in begin - #foc + (#r + #ctx) ≡˘⟨ +-assoc #foc #r #ctx ⟩ + #foc + (#r + #ctx) ≡⟨ +-assoc #foc #r #ctx ⟨ #leaves (mkZipper ctx (node foc m r)) ≡⟨ helper ctx (node foc m r) ⟩ BT.#leaves (toTree (mkZipper cs foc)) ∎ diff --git a/src/Data/Tree/Rose/Properties.agda b/src/Data/Tree/Rose/Properties.agda index 749bdb5748..054103fdd1 100644 --- a/src/Data/Tree/Rose/Properties.agda +++ b/src/Data/Tree/Rose/Properties.agda @@ -39,7 +39,7 @@ map-∘ f g (node a ts) = cong (node (g (f a))) $ begin depth-map : (f : A → B) (t : Rose A i) → depth {i = i} (map {i = i} f t) ≡ depth {i = i} t depth-map f (node a ts) = cong (suc ∘′ max 0) $ begin - List.map depth (List.map (map f) ts) ≡˘⟨ Listₚ.map-∘ ts ⟩ + List.map depth (List.map (map f) ts) ≡⟨ Listₚ.map-∘ ts ⟨ List.map (depth ∘′ map f) ts ≡⟨ Listₚ.map-cong (depth-map f) ts ⟩ List.map depth ts ∎ @@ -49,7 +49,7 @@ depth-map f (node a ts) = cong (suc ∘′ max 0) $ begin foldr-map : (f : A → B) (n : B → List C → C) (ts : Rose A i) → foldr {i = i} n (map {i = i} f ts) ≡ foldr {i = i} (n ∘′ f) ts foldr-map f n (node a ts) = cong (n (f a)) $ begin - List.map (foldr n) (List.map (map f) ts) ≡˘⟨ Listₚ.map-∘ ts ⟩ + List.map (foldr n) (List.map (map f) ts) ≡⟨ Listₚ.map-∘ ts ⟨ List.map (foldr n ∘′ map f) ts ≡⟨ Listₚ.map-cong (foldr-map f n) ts ⟩ List.map (foldr (n ∘′ f)) ts ∎ diff --git a/src/Data/Vec/Functional/Properties.agda b/src/Data/Vec/Functional/Properties.agda index 93f95b423b..d1d5f8c628 100644 --- a/src/Data/Vec/Functional/Properties.agda +++ b/src/Data/Vec/Functional/Properties.agda @@ -181,20 +181,20 @@ module _ {ys ys′ : Vector A m} where ... | yes i Date: Thu, 19 Oct 2023 10:35:34 +0900 Subject: [PATCH 35/57] Restore 'return' as an alias for 'pure' (#2164) --- CHANGELOG.md | 3 +-- src/Effect/Applicative.agda | 4 ++++ src/IO/Primitive.agda | 17 +++++++++++++---- 3 files changed, 18 insertions(+), 6 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 4d6a8c2ccd..3721f2b454 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -774,8 +774,7 @@ Non-backwards compatible changes 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) + `RawApplicative` + `_>>=_`. This reorganisation means in particular that the functor/applicative of a monad are not computed using `_>>=_`. This may break proofs. diff --git a/src/Effect/Applicative.agda b/src/Effect/Applicative.agda index 729966752c..a2263c6e0a 100644 --- a/src/Effect/Applicative.agda +++ b/src/Effect/Applicative.agda @@ -53,6 +53,10 @@ record RawApplicative (F : Set f → Set g) : Set (suc f ⊔ g) where zip : F A → F B → F (A × B) zip = zipWith _,_ + -- Haskell-style alternative name for pure + return : A → F A + return = pure + -- backwards compatibility: unicode variants _⊛_ : F (A → B) → F A → F B _⊛_ = _<*>_ diff --git a/src/IO/Primitive.agda b/src/IO/Primitive.agda index 220e862fb4..8014a2bd7c 100644 --- a/src/IO/Primitive.agda +++ b/src/IO/Primitive.agda @@ -10,20 +10,29 @@ module IO.Primitive where -open import Agda.Builtin.IO +open import Level using (Level) +private + variable + a : Level + A B : Set a ------------------------------------------------------------------------ -- The IO monad -open import Agda.Builtin.IO public using (IO) +open import Agda.Builtin.IO public + using (IO) infixl 1 _>>=_ postulate - pure : ∀ {a} {A : Set a} → A → IO A - _>>=_ : ∀ {a b} {A : Set a} {B : Set b} → IO A → (A → IO B) → IO B + pure : A → IO A + _>>=_ : IO A → (A → IO B) → IO B {-# COMPILE GHC pure = \_ _ -> return #-} {-# COMPILE GHC _>>=_ = \_ _ _ _ -> (>>=) #-} {-# COMPILE UHC pure = \_ _ x -> UHC.Agda.Builtins.primReturn x #-} {-# COMPILE UHC _>>=_ = \_ _ _ _ x y -> UHC.Agda.Builtins.primBind x y #-} + +-- Haskell-style alternative syntax +return : A → IO A +return = pure From ce4bd12157250f1659cc09282bafac3d50893d5e Mon Sep 17 00:00:00 2001 From: "G. Allais" Date: Thu, 19 Oct 2023 02:35:54 +0100 Subject: [PATCH 36/57] [ fix #2153 ] Properly re-export specialised combinators (#2161) --- CHANGELOG.md | 3 +++ src/Data/List/Membership/Setoid.agda | 5 +++-- 2 files changed, 6 insertions(+), 2 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 3721f2b454..e05088b863 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -33,6 +33,9 @@ Bug-fixes in `Function.Construct.Composition` had their arguments in the wrong order. They have been flipped so they can actually be used as a composition operator. +* The operations `_∷=_` and `_─_` exported from `Data.List.Membership.Setoid` + had an extraneous `{A : Set a}` parameter. This has been removed. + * The combinators `_≃⟨_⟩_` and `_≃˘⟨_⟩_` in `Data.Rational.Properties.≤-Reasoning` now correctly accepts proofs of type `_≃_` instead of the previous proofs of type `_≡_`. diff --git a/src/Data/List/Membership/Setoid.agda b/src/Data/List/Membership/Setoid.agda index 1b8a67fb7d..1a19169992 100644 --- a/src/Data/List/Membership/Setoid.agda +++ b/src/Data/List/Membership/Setoid.agda @@ -13,7 +13,7 @@ module Data.List.Membership.Setoid {c ℓ} (S : Setoid c ℓ) where open import Function.Base using (_∘_; id; flip) open import Data.List.Base as List using (List; []; _∷_; length; lookup) -open import Data.List.Relation.Unary.Any +open import Data.List.Relation.Unary.Any as Any using (Any; index; map; here; there) open import Data.Product.Base as Prod using (∃; _×_; _,_) open import Relation.Unary using (Pred) @@ -35,7 +35,8 @@ x ∉ xs = ¬ x ∈ xs ------------------------------------------------------------------------ -- Operations -open Data.List.Relation.Unary.Any using (_∷=_; _─_) public +_∷=_ = Any._∷=_ {A = A} +_─_ = Any._─_ {A = A} mapWith∈ : ∀ {b} {B : Set b} (xs : List A) → (∀ {x} → x ∈ xs → B) → List B From c5a36939b6e5aafb2ce12bd59af3153a7da4daa0 Mon Sep 17 00:00:00 2001 From: James Wood Date: Thu, 19 Oct 2023 03:39:33 +0100 Subject: [PATCH 37/57] Connect `LeftInverse` with (`Split`)`Surjection` (#2054) --- CHANGELOG.md | 13 +++++++++ src/Function/Bundles.agda | 55 +++++++++++++++++++++++++++++++++++- src/Function/Structures.agda | 20 +++++++++++++ 3 files changed, 87 insertions(+), 1 deletion(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index e05088b863..998c63a43c 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -3238,6 +3238,13 @@ Additions to existing modules evalState : State s a → s → a execState : State s a → s → s ``` + +* Added new proofs and definitions in `Function.Bundles`: + ```agda + LeftInverse.isSplitSurjection : LeftInverse → IsSplitSurjection to + LeftInverse.surjection : LeftInverse → Surjection + SplitSurjection = LeftInverse + ``` * Added new application operator synonym to `Function.Bundles`: ```agda @@ -3260,6 +3267,12 @@ Additions to existing modules _!|>_ : (a : A) → (∀ a → B a) → B a _!|>′_ : A → (A → B) → B ``` + +* Added new proof and record in `Function.Structures`: + ```agda + IsLeftInverse.isSurjection : IsLeftInverse to from → IsSurjection to + record IsSplitSurjection (f : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) + ``` * Added new definition to the `Surjection` module in `Function.Related.Surjection`: ``` diff --git a/src/Function/Bundles.agda b/src/Function/Bundles.agda index 0aef27fc57..9e18c30fe5 100644 --- a/src/Function/Bundles.agda +++ b/src/Function/Bundles.agda @@ -218,7 +218,7 @@ module _ (From : Setoid a ℓ₁) (To : Setoid b ℓ₂) where } open IsLeftInverse isLeftInverse public - using (module Eq₁; module Eq₂; strictlyInverseˡ) + using (module Eq₁; module Eq₂; strictlyInverseˡ; isSurjection) equivalence : Equivalence equivalence = record @@ -226,6 +226,20 @@ module _ (From : Setoid a ℓ₁) (To : Setoid b ℓ₂) where ; from-cong = from-cong } + isSplitSurjection : IsSplitSurjection to + isSplitSurjection = record + { from = from + ; isLeftInverse = isLeftInverse + } + + surjection : Surjection From To + surjection = record + { to = to + ; cong = to-cong + ; surjective = λ y → from y , inverseˡ + } + + record RightInverse : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) where field @@ -346,6 +360,45 @@ module _ (From : Setoid a ℓ₁) (To : Setoid b ℓ₂) where ; from₂-cong = from₂-cong } +------------------------------------------------------------------------ +-- Other + + -- A left inverse is also known as a “split surjection”. + -- + -- As the name implies, a split surjection is a special kind of + -- surjection where the witness generated in the domain in the + -- function for elements `x₁` and `x₂` are equal if `x₁ ≈ x₂` . + -- + -- The difference is the `from-cong` law --- generally, the section + -- (called `Surjection.to⁻` or `SplitSurjection.from`) of a surjection + -- need not respect equality, whereas it must in a split surjection. + -- + -- The two notions coincide when the equivalence relation on `B` is + -- propositional equality (because all functions respect propositional + -- equality). + -- + -- For further background on (split) surjections, one may consult any + -- general mathematical references which work without the principle + -- of choice. For example: + -- + -- https://ncatlab.org/nlab/show/split+epimorphism. + -- + -- The connection to set-theoretic notions with the same names is + -- justified by the setoid type theory/homotopy type theory + -- observation/definition that (∃x : A. P) = ∥ Σx : A. P ∥ --- i.e., + -- we can read set-theoretic ∃ as squashed/propositionally truncated Σ. + -- + -- We see working with setoids as working in the MLTT model of a setoid + -- type theory, in which ∥ X ∥ is interpreted as the setoid with carrier + -- set X and the equivalence relation that relates all elements. + -- All maps into ∥ X ∥ respect equality, so in the idiomatic definitions + -- here, we drop the corresponding trivial `cong` field completely. + + SplitSurjection : Set _ + SplitSurjection = LeftInverse + + module SplitSurjection (splitSurjection : SplitSurjection) = + LeftInverse splitSurjection ------------------------------------------------------------------------ -- Bundles specialised for propositional equality diff --git a/src/Function/Structures.agda b/src/Function/Structures.agda index a1cc0d6ceb..baf33dddf0 100644 --- a/src/Function/Structures.agda +++ b/src/Function/Structures.agda @@ -107,6 +107,12 @@ record IsLeftInverse (to : A → B) (from : B → A) : Set (a ⊔ b ⊔ ℓ₁ strictlyInverseˡ : StrictlyInverseˡ _≈₂_ to from strictlyInverseˡ x = inverseˡ Eq₁.refl + isSurjection : IsSurjection to + isSurjection = record + { isCongruent = isCongruent + ; surjective = λ y → from y , inverseˡ + } + record IsRightInverse (to : A → B) (from : B → A) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) where field @@ -168,3 +174,17 @@ record IsBiInverse open IsCongruent to-isCongruent public renaming (cong to to-cong) + + +------------------------------------------------------------------------ +-- Other +------------------------------------------------------------------------ + +-- See the comment on `SplitSurjection` in `Function.Bundles` for an +-- explanation of (split) surjections. +record IsSplitSurjection (f : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) where + field + from : B → A + isLeftInverse : IsLeftInverse f from + + open IsLeftInverse isLeftInverse public From eb5f30854bc8f8af5e6b7038f15656edb493f0cf Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Thu, 19 Oct 2023 11:39:49 +0900 Subject: [PATCH 38/57] Added remaining flipped and negated relations to binary relation bundles (#2162) --- CHANGELOG.md | 30 ++++--- README/Design/Hierarchies.agda | 37 ++++++++ .../Lattice/Properties/Semilattice.agda | 4 +- .../Product/Relation/Binary/Lex/Strict.agda | 2 +- src/Relation/Binary/Bundles.agda | 88 +++++++++++++------ .../Binary/Properties/DecTotalOrder.agda | 3 +- src/Relation/Binary/Properties/Poset.agda | 5 -- .../Binary/Properties/TotalOrder.agda | 5 +- src/Relation/Binary/Structures.agda | 33 ++++--- 9 files changed, 144 insertions(+), 63 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 998c63a43c..4b67ec1849 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -852,18 +852,24 @@ Non-backwards compatible changes IO.Instances ``` -### (Issue #2096) Introduction of flipped relation symbol for `Relation.Binary.Bundles.Preorder` - -* Previously, the relation symbol `_∼_` was (notationally) symmetric, so that its - converse relation could only be discussed *semantically* in terms of `flip _∼_` - in `Relation.Binary.Properties.Preorder`, `Relation.Binary.Construct.Flip.{Ord|EqAndOrd}` - -* Now, the symbol `_∼_` has been renamed to a new symbol `_≲_`, with `_≳_` - introduced as a definition in `Relation.Binary.Bundles.Preorder` whose properties - in `Relation.Binary.Properties.Preorder` now refer to it. Partial backwards compatible - has been achieved by redeclaring a deprecated version of the old name in the record. - Therefore, only _declarations_ of `PartialOrder` records will need their field names - updating. +### (Issue #2096) Introduction of flipped and negated relation symbols to bundles in `Relation.Binary.Bundles` + +* Previously, bundles such as `Preorder`, `Poset`, `TotalOrder` etc. did not have the flipped + and negated versions of the operators exposed. In some cases they could obtained by opening the + relevant `Relation.Binary.Properties.X` file but usually they had to be redefined every time. + +* To fix this, these bundles now all export all 4 versions of the operator: normal, converse, negated, + converse-negated. Accordingly they are no longer exported from the corresponding `Properties` file. + +* To make this work for `Preorder`, it was necessary to change the name of the relation symbol. + Previously, the symbol was `_∼_` which is (notationally) symmetric, so that its + converse relation could only be discussed *semantically* in terms of `flip _∼_`. + +* Now, the `Preorder` record field `_∼_` has been renamed to `_≲_`, with `_≳_` + introduced as a definition in `Relation.Binary.Bundles.Preorder`. + Partial backwards compatible has been achieved by redeclaring a deprecated version + of the old symbol in the record. Therefore, only _declarations_ of `PartialOrder` records will + need their field names updating. ### (Issue #1214) Reorganisation of the introduction of negated relation symbols under `Relation.Binary` diff --git a/README/Design/Hierarchies.agda b/README/Design/Hierarchies.agda index 09d5504a91..bc9220ba23 100644 --- a/README/Design/Hierarchies.agda +++ b/README/Design/Hierarchies.agda @@ -249,6 +249,43 @@ record Semigroup : Set (suc (a ⊔ ℓ)) where -- differently, as a function with an unknown domain an codomain is -- of little use. +------------------------- +-- Bundle re-exporting -- +------------------------- + +-- In general ensuring that bundles re-export everything in their +-- sub-bundles can get a little tricky. + +-- Imagine we have the following general scenario where bundle A is a +-- direct refinement of bundle C (i.e. the record `IsA` has a `IsC` field) +-- but is also morally a refinement of bundles B and D. + +-- Structures Bundles +-- ========== ======= +-- IsA A +-- / || \ / || \ +-- IsB IsC IsD B C D + +-- The procedure for re-exports in the bundles is as follows: + +-- 1. `open IsA isA public using (IsC, M)` where `M` is everything +-- exported by `IsA` that is not exported by `IsC`. + +-- 2. Construct `c : C` via the `isC` obtained in step 1. + +-- 3. `open C c public hiding (N)` where `N` is the list of fields +-- shared by both `A` and `C`. + +-- 4. Construct `b : B` via the `isB` obtained in step 1. + +-- 5. `open B b public using (O)` where `O` is everything exported +-- by `B` but not exported by `IsA`. + +-- 6. Construct `d : D` via the `isC` obtained in step 1. + +-- 7. `open D d public using (P)` where `P` is everything exported +-- by `D` but not exported by `IsA` + ------------------------------------------------------------------------ -- Other hierarchy modules ------------------------------------------------------------------------ diff --git a/src/Algebra/Lattice/Properties/Semilattice.agda b/src/Algebra/Lattice/Properties/Semilattice.agda index 38554aedfa..7b8135ced7 100644 --- a/src/Algebra/Lattice/Properties/Semilattice.agda +++ b/src/Algebra/Lattice/Properties/Semilattice.agda @@ -27,8 +27,8 @@ import Relation.Binary.Construct.NaturalOrder.Left _≈_ _∧_ poset : Poset c ℓ ℓ poset = LeftNaturalOrder.poset isSemilattice -open Poset poset using (_≤_; isPartialOrder) -open PosetProperties poset using (_≥_; ≥-isPartialOrder) +open Poset poset using (_≤_; _≥_; isPartialOrder) +open PosetProperties poset using (≥-isPartialOrder) ------------------------------------------------------------------------ -- Every algebraic semilattice can be turned into an order-theoretic one. diff --git a/src/Data/Product/Relation/Binary/Lex/Strict.agda b/src/Data/Product/Relation/Binary/Lex/Strict.agda index 42d139893d..39d86602f2 100644 --- a/src/Data/Product/Relation/Binary/Lex/Strict.agda +++ b/src/Data/Product/Relation/Binary/Lex/Strict.agda @@ -240,7 +240,7 @@ module _ {_≈₁_ : Rel A ℓ₁} {_<₁_ : Rel A ℓ₂} (isEquivalence pre₁) (isEquivalence pre₂) ; reflexive = ×-reflexive _≈₁_ _<₁_ _<₂_ (reflexive pre₂) ; trans = ×-transitive {_<₂_ = _<₂_} - (isEquivalence pre₁) (∼-resp-≈ pre₁) + (isEquivalence pre₁) (≲-resp-≈ pre₁) (trans pre₁) (trans pre₂) } where open IsPreorder diff --git a/src/Relation/Binary/Bundles.agda b/src/Relation/Binary/Bundles.agda index eb7975f376..641803f8ef 100644 --- a/src/Relation/Binary/Bundles.agda +++ b/src/Relation/Binary/Bundles.agda @@ -43,13 +43,15 @@ record Setoid c ℓ : Set (suc (c ⊔ ℓ)) where isEquivalence : IsEquivalence _≈_ open IsEquivalence isEquivalence public + using (refl; reflexive; isPartialEquivalence) partialSetoid : PartialSetoid c ℓ partialSetoid = record { isPartialEquivalence = isPartialEquivalence } - open PartialSetoid partialSetoid public using (_≉_) + open PartialSetoid partialSetoid public + hiding (Carrier; _≈_; isPartialEquivalence) record DecSetoid c ℓ : Set (suc (c ⊔ ℓ)) where @@ -60,14 +62,15 @@ record DecSetoid c ℓ : Set (suc (c ⊔ ℓ)) where isDecEquivalence : IsDecEquivalence _≈_ open IsDecEquivalence isDecEquivalence public + using (_≟_; isEquivalence) setoid : Setoid c ℓ setoid = record { isEquivalence = isEquivalence } - open Setoid setoid public using (partialSetoid; _≉_) - + open Setoid setoid public + hiding (Carrier; _≈_; isEquivalence) ------------------------------------------------------------------------ -- Preorders @@ -99,6 +102,9 @@ record Preorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≳_ _≳_ = flip _≲_ + infix 4 _⋧_ + _⋧_ = flip _⋦_ + -- Deprecated. infix 4 _∼_ _∼_ = _≲_ @@ -117,14 +123,15 @@ record TotalPreorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where isTotalPreorder : IsTotalPreorder _≈_ _≲_ open IsTotalPreorder isTotalPreorder public - hiding (module Eq) + using (total; isPreorder) preorder : Preorder c ℓ₁ ℓ₂ - preorder = record { isPreorder = isPreorder } + preorder = record + { isPreorder = isPreorder + } open Preorder preorder public - using (module Eq; _≳_; _⋦_) - + hiding (Carrier; _≈_; _≲_; isPreorder) ------------------------------------------------------------------------ -- Partial orders @@ -139,7 +146,7 @@ record Poset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where isPartialOrder : IsPartialOrder _≈_ _≤_ open IsPartialOrder isPartialOrder public - hiding (module Eq) + using (antisym; isPreorder) preorder : Preorder c ℓ₁ ℓ₂ preorder = record @@ -147,8 +154,13 @@ record Poset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where } open Preorder preorder public - using (module Eq) - renaming (_⋦_ to _≰_) + hiding (Carrier; _≈_; _≲_; isPreorder) + renaming + ( _⋦_ to _≰_; _≳_ to _≥_; _⋧_ to _≱_ + ; ≲-respˡ-≈ to ≤-respˡ-≈ + ; ≲-respʳ-≈ to ≤-respʳ-≈ + ; ≲-resp-≈ to ≤-resp-≈ + ) record DecPoset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where @@ -159,9 +171,10 @@ record DecPoset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where _≤_ : Rel Carrier ℓ₂ isDecPartialOrder : IsDecPartialOrder _≈_ _≤_ - private - module DPO = IsDecPartialOrder isDecPartialOrder - open DPO public hiding (module Eq) + private module DPO = IsDecPartialOrder isDecPartialOrder + + open DPO public + using (_≟_; _≤?_; isPartialOrder) poset : Poset c ℓ₁ ℓ₂ poset = record @@ -169,7 +182,7 @@ record DecPoset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where } open Poset poset public - using (preorder; _≰_) + hiding (Carrier; _≈_; _≤_; isPartialOrder; module Eq) module Eq where decSetoid : DecSetoid c ℓ₁ @@ -203,6 +216,12 @@ record StrictPartialOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) _≮_ : Rel Carrier _ x ≮ y = ¬ (x < y) + infix 4 _>_ + _>_ = flip _<_ + + infix 4 _≯_ + _≯_ = flip _≮_ + record DecStrictPartialOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where infix 4 _≈_ _<_ @@ -212,9 +231,10 @@ record DecStrictPartialOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂ _<_ : Rel Carrier ℓ₂ isDecStrictPartialOrder : IsDecStrictPartialOrder _≈_ _<_ - private - module DSPO = IsDecStrictPartialOrder isDecStrictPartialOrder - open DSPO public hiding (module Eq) + private module DSPO = IsDecStrictPartialOrder isDecStrictPartialOrder + + open DSPO public + using (_ Date: Thu, 19 Oct 2023 19:19:07 +0900 Subject: [PATCH 39/57] Tidy up CHANGELOG in preparation for release candidate (#2165) --- CHANGELOG.md | 3151 ++++++++++++++++++++++++-------------------------- 1 file changed, 1483 insertions(+), 1668 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 4b67ec1849..29113663c0 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1,259 +1,156 @@ -Version 2.0-dev +Version 2.0-rc1 =============== -The library has been tested using Agda 2.6.3. +The library has been tested using Agda 2.6.4. + +NOTE: Version `2.0` contains various breaking changes and is not backwards +compatible with code written with version `1.X` of the library. 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 +* 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. +* Significant tightening of `import` statements internally in the library, + drastically decreasing the dependencies and hence load time of many key + modules. + +* 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. Bug-fixes --------- -* The combinators `_⟶-∘_`, `_↣-∘_`, `_↠-∘_`, `_⤖-∘_`, `_⇔-∘_`, `_↩-∘_`, `_↪-∘_`, `_↔-∘_` - in `Function.Construct.Composition` had their arguments in the wrong order. They have - been flipped so they can actually be used as a composition operator. +* In `Algebra.Definitions.RawSemiring` the record `Prime` did not + enforce that the number was not divisible by `1#`. To fix this + `p∤1 : p ∤ 1#` hsa been added as a field. -* The operations `_∷=_` and `_─_` exported from `Data.List.Membership.Setoid` +* 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.Fin.Properties` the `i` argument to `opposite-suc` was implicit + but could not be inferred in general. It has been made explicit. + +* In `Data.List.Membership.Setoid` the operations `_∷=_` and `_─_` had an extraneous `{A : Set a}` parameter. This has been removed. -* The combinators `_≃⟨_⟩_` and `_≃˘⟨_⟩_` in `Data.Rational.Properties.≤-Reasoning` - now correctly accepts proofs of type `_≃_` instead of the previous proofs of type `_≡_`. +* In `Data.List.Relation.Ternary.Appending.Setoid` the constructors + were re-exported in their full generality which lead to unsolved meta + variables at their use sites. Now versions of the constructors specialised + to use the setoid's carrier set are re-exported. + +* In `Data.Nat.DivMod` the parameter `o` in the proof `/-monoˡ-≤` was + implicit but not inferrable. It has been changed to be explicit. -* 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) - infixr 4 _,_ (Data.Refinement) - infixr 4 _,_ (Data.Container.Relation.Binary.Pointwise) - infixr 4 _,_ (Data.Tree.AVL.Value) - infixr 4 _,_ (Foreign.Haskell.Pair) - infixr 4 _,_ (Reflection.AnnotatedAST) - infixr 4 _,_ (Reflection.AST.Traversal) - infixl 6.5 _P′_ _P_ _C′_ _C_ (Data.Nat.Combinatorics.Base) - infixl 1 _>>=-cong_ _≡->>=-cong_ (Effect.Monad.Partiality) - infixl 1 _?>=′_ (Effect.Monad.Predicate) - infixl 1 _>>=-cong_ _>>=-congP_ (Effect.Monad.Partiality.All) - infix 4 _∈FV_ (Reflection.AST.DeBruijn) - infixr 9 _;_ (Relation.Binary.Construct.Composition) - infixl 6 _+²_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) - infixr -1 _atₛ_ (Relation.Binary.Indexed.Heterogeneous.Construct.At) - infixr -1 _atₛ_ (Relation.Binary.Indexed.Homogeneous.Construct.At) - infix 4 _∈_ _∉_ (Relation.Unary.Indexed) - infixr 9 _⍮_ (Relation.Unary.PredicateTransformer) - infix 8 ∼_ (Relation.Unary.PredicateTransformer) - infix 2 _×?_ _⊙?_ (Relation.Unary.Properties) - infix 10 _~? (Relation.Unary.Properties) - infixr 1 _⊎?_ (Relation.Unary.Properties) - infixr 7 _∩?_ (Relation.Unary.Properties) - infixr 6 _∪?_ (Relation.Unary.Properties) - infixl 6 _`⊜_ (Tactic.RingSolver) - infix 8 ⊝_ (Tactic.RingSolver.Core.Expression) - infix 4 _∈ᴿ?_ _∉ᴿ?_ _∈?ε _∈?[_] _∈?[^_] (Text.Regex.Properties) - infix 4 _∈?_ _∉?_ (Text.Regex.Derivative.Brzozowski) - infix 4 _∈_ _∉_ _∈?_ _∉?_ (Text.Regex.String.Unsafe) - ``` +* In `Data.Nat.DivMod` the parameter `m` in the proof `+-distrib-/-∣ʳ` was + implicit but not inferrable, while `n` is explicit but inferrable. + They have been to explicit and implicit respectively. -* 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 `Data.Nat.GeneralisedArithmetic` the `s` and `z` arguments to the + following functions were implicit but no inferrable: + `fold-+`, `fold-k`, `fold-*`, `fold-pull`. They have been made explicit. -* In `/-monoˡ-≤` in `Data.Nat.DivMod` the parameter `o` was implicit but not inferrable. - It has been changed to be explicit. +* In `Data.Rational(.Unnormalised).Properties` the module `≤-Reasoning` + exported equality combinators using the generic setoid symbol `_≈_`. They + have been renamed to use the same `_≃_` symbol used for non-propositional + equality over `Rational`s, i.e. + ```agda + step-≈ ↦ step-≃ + step-≈˘ ↦ step-≃˘ + ``` + with corresponding associated syntax: + ```agda + _≈⟨_⟩_ ↦ _≃⟨_⟩_ + _≈⟨_⟨_ ↦ _≃⟨_⟨_ + ``` -* 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.Construct.Composition` the combinators + `_⟶-∘_`, `_↣-∘_`, `_↠-∘_`, `_⤖-∘_`, `_⇔-∘_`, `_↩-∘_`, `_↪-∘_`, `_↔-∘_` + had their arguments in the wrong order. They have been flipped so they can + actually be used as a composition operator. * 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. +* 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. Non-backwards compatible changes -------------------------------- -#### Removed deprecated names +### Removed deprecated names -* All modules and names that were deprecated in v1.2 and before have been removed. +* All modules and names that were deprecated in `v1.2` and before have + been removed. -#### Moved `Codata` modules to `Codata.Sized` +### Changes to `LeftCancellative` and `RightCancellative` in `Algebra.Definitions` -* 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`. +* 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. -#### Moved `Category` modules to `Effect` +* 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. -* 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. +* 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` -### Improvements to pretty printing and regexes + - `RightCancellative _•_` + - From: `∀ {x} y z → (y • x) ≈ (z • x) → y ≈ z` + - To: `∀ x y z → (y • x) ≈ (z • x) → y ≈ z` -* 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 + - `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` - 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 + - `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` - main = do - -- you can now use the same pretty with different widths: - putStrLn $ render 40 pretty - putStrLn $ render 80 pretty - ``` +* 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` -* 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. +### Changes to ring structures in `Algebra` - 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". +* 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*` -### Refactoring of algebraic lattice hierarchy +### Refactoring of lattices in `Algebra.Structures/Bundles` hierarchy * In order to improve modularity and consistency with `Relation.Binary.Lattice`, the structures & bundles for `Semilattice`, `Lattice`, `DistributiveLattice` @@ -265,9 +162,9 @@ Non-backwards compatible changes `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: +* The definition of `IsDistributiveLattice` and `IsBooleanAlgebra` have changed so + that they are no longer right-biased which hindered compositionality. + More concretely, `IsDistributiveLattice` now has fields: ```agda ∨-distrib-∧ : ∨ DistributesOver ∧ ∧-distrib-∨ : ∧ DistributesOver ∨ @@ -277,7 +174,7 @@ Non-backwards compatible changes ∨-distribʳ-∧ : ∨ DistributesOverʳ ∧ ``` and `IsBooleanAlgebra` now has fields: - ``` + ```agda ∨-complement : Inverse ⊤ ¬ ∨ ∧-complement : Inverse ⊥ ¬ ∧ ``` @@ -310,140 +207,326 @@ Non-backwards compatible changes * 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`. +* Finally, the following auxiliary files have been moved: + ```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 + ``` -* 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`. +#### Changes to `Algebra.Morphism.Structures` -* In `Function.Definitions` the types of `Surjective`, `Injective` and `Surjective` - have been changed from: +* Previously the record definitions: ``` - Surjective f = ∀ y → ∃ λ x → f x ≈₂ y - Inverseˡ f g = ∀ y → f (g y) ≈₂ y - Inverseʳ f g = ∀ x → g (f x) ≈₁ x + IsNearSemiringHomomorphism + IsSemiringHomomorphism + IsRingHomomorphism ``` - to: + all had two separate proofs of `IsRelHomomorphism` within them. + +* To fix this they have all been redefined to build up from `IsMonoidHomomorphism`, + `IsNearSemiringHomomorphism`, and `IsSemiringHomomorphism` respectively, + adding a single property at each step. + +* Similarly, `IsLatticeHomomorphism` is now built as + `IsRelHomomorphism` along with proofs that `_∧_` and `_∨_` are homomorphic. + +* Finally `⁻¹-homo` in `IsRingHomomorphism` has been renamed to `-‿homo`. + +#### Renamed `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. + +* Full list of moved modules: + ```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 ``` - 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 + +* Full list of 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 ``` - 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)`. +### Refactoring of the unindexed Functor/Applicative/Monad hierarchy in `Effect` + +* 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` + `_>>=_`. + 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. -#### Proofs of non-zeroness/positivity/negativity as instance arguments + +#### 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 `Codata.Sized`, e.g. `Codata.Stream` + has become `Codata.Sized.Stream`. + +### New proof-irrelevant for empty type in `Data.Empty` + +* 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 the following two definitions have been modified: + + + 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 + ``` + +### Deprecation of `_≺_` in `Data.Fin.Base` + +* In `Data.Fin.Base` the relation `_≺_` and its single constructor `_≻toℕ_` + have been deprecated in favour of their extensional equivalent `_<_` + but omitting the inversion principle which pattern matching on `_≻toℕ_` + would achieve; this instead is proxied by the property `Data.Fin.Properties.toℕ<`. + +* Consequently in `Data.Fin.Induction`: + ``` + ≺-Rec + ≺-wellFounded + ≺-recBuilder + ≺-rec + ``` + these functions are also deprecated. + +* Likewise in `Data.Fin.Properties` the proofs `≺⇒<′` and `<′⇒≺` have been deprecated + in favour of their proxy counterparts `<⇒<′` and `<′⇒<`. + +### Standardisation of `insertAt`/`updateAt`/`removeAt` in `Data.List`/`Data.Vec` + +* Previously, the names and argument order of index-based insertion, update and removal functions for + various types of lists and vectors were inconsistent. + +* To fix this the names have all been standardised to `insertAt`/`updateAt`/`removeAt`. + +* Correspondingly the following changes have occurred: + +* In `Data.List.Base` the following have been added: + ```agda + insertAt : (xs : List A) → Fin (suc (length xs)) → A → List A + updateAt : (xs : List A) → Fin (length xs) → (A → A) → List A + removeAt : (xs : List A) → Fin (length xs) → List A + ``` + and the following has been deprecated + ``` + _─_ ↦ removeAt + ``` + +* In `Data.Vec.Base`: + ```agda + insert ↦ insertAt + remove ↦ removeAt + + updateAt : Fin n → (A → A) → Vec A n → Vec A n + ↦ + updateAt : Vec A n → Fin n → (A → A) → Vec A n + ``` + +* In `Data.Vec.Functional`: + ```agda + remove : Fin (suc n) → Vector A (suc n) → Vector A n + ↦ + removeAt : Vector A (suc n) → Fin (suc n) → Vector A n + + updateAt : Fin n → (A → A) → Vector A n → Vector A n + ↦ + updateAt : Vector A n → Fin n → (A → A) → Vector A n + ``` + +* The old names (and the names of all proofs about these functions) have been deprecated appropriately. + +#### Standardisation of `lookup` in `Data.(List/Vec/...)` + +* All the types of `lookup` functions (and variants) in the following modules + have been changed to match the argument convention adopted in the `List` module (i.e. + `lookup` takes its container first and the index, whose type may depend on the + container value, second): + ``` + 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 + ``` + +* To accomodate this in in `Data.Vec.Relation.Unary.Linked.Properties` + and `Codata.Guarded.Stream.Relation.Binary.Pointwise`, the proofs + called `lookup` have been renamed `lookup⁺`. + +#### Changes to `Data.(Nat/Integer/Rational)` proofs of `NonZero`/`Positive`/`Negative` to 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 + 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. + See the mailing list discussion for a fuller explanation of the motivation and implementation. -* For example, whereas the type signature of division used to be: - ``` +* For example, whereas the type of division over `ℕ` used to be: + ```agda _/_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} → ℕ ``` it is now: - ``` + ```agda _/_ : (dividend divisor : ℕ) .{{_ : NonZero divisor}} → ℕ ``` - which means that as long as an instance of `NonZero n` is in scope then you can write + +* This 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_` @@ -454,62 +537,120 @@ Non-backwards compatible changes * 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. + + 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. + ```agda + 0/n≡0 : 0 / suc n ≡ 0 + ``` + + 2. *Add an instance argument parameter* - 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. + ```agda + 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. + ```agda + 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. + This made the proofs extremely difficult to use if your term wasn't in the form `suc n`. 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. + 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`: +* Finally, in `Data.Rational.Unnormalised.Base` the definition of `_≢0` is now + redundent and so has been removed. Additionally the following proofs about it have + also 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 +### Changes to the definition of `_≤″_` in `Data.Nat.Base` (issue #1919) + +* The definition of `_≤″_` was previously: + ```agda + record _≤″_ (m n : ℕ) : Set where + constructor less-than-or-equal + field + {k} : ℕ + proof : m + k ≡ n + ``` + which introduced a spurious additional definition, when this is in fact, modulo + field names and implicit/explicit qualifiers, equivalent to the definition of left- + divisibility, `_∣ˡ_` for the `RawMagma` structure of `_+_`. + +* Since the addition of raw bundles to `Data.X.Base`, this definition can now be + made directly. Accordingly, the definition has been changed to: + ```agda + _≤″_ : (m n : ℕ) → Set + _≤″_ = _∣ˡ_ +-rawMagma + + pattern less-than-or-equal {k} prf = k , prf + ``` + +* Knock-on consequences include the need to retain the old constructor + name, now introduced as a pattern synonym, and introduction of (a function + equivalent to) the former field name/projection function `proof` as + `≤″-proof` in `Data.Nat.Properties`. + +### Changes to definition of `Prime` in `Data.Nat.Primality` + +* The definition of `Prime` 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 + ``` + +### Changes to operation reduction behaviour in `Data.Rational(.Unnormalised)` * Currently arithmetic expressions involving rationals (both normalised and unnormalised) undergo disastrous exponential normalisation. For example, @@ -530,118 +671,114 @@ Non-backwards compatible changes 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: +* As a consequence of this, some proofs that relied either on this reduction behaviour + or on eta-equality may no longer type-check. + +* 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. * Similarly, in order to prevent reduction, the equality `_≃_` in `Data.Rational.Base` - has been made into a data type with the single constructor `*≡*`. The destructor `drop-*≡*` - has been added to `Data.Rational.Properties`. + has been made into a data type with the single constructor `*≡*`. The destructor + `drop-*≡*` has been added to `Data.Rational.Properties`. -### 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. +#### Deprecation of old `Function` hierarchy -* 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. +* The new `Function` hierarchy was introduced in `v1.2` which follows + the same `Core`/`Definitions`/`Bundles`/`Structures` as all the other + hierarchies in the library. + +* At the time the old function hierarchy in: + ``` + Function.Equality + Function.Injection + Function.Surjection + Function.Bijection + Function.LeftInverse + Function.Inverse + Function.HalfAdjointEquivalence + Function.Related + ``` + was unofficially deprecated, but could not be officially deprecated because + of it's widespread use in the rest of the library. + +* Now, the old hierarchy modules have all been officially deprecated. All + uses of them in the rest of the library have been switched to use the + new hierarachy. + +* The latter is unfortunately a relatively big breaking change, but was judged + to be unavoidable given how widely used the old hierarchy was. -* 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` +#### Changes to the new `Function` hierarchy - - `RightCancellative _•_` - - From: `∀ {x} y z → (y • x) ≈ (z • x) → y ≈ z` - - To: `∀ x y z → (y • x) ≈ (z • x) → y ≈ z` +* In `Function.Bundles` the names of the fields in the records have been + changed from `f`, `g`, `cong₁` and `cong₂` to `to`, `from`, `to-cong`, `from-cong`. - - `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` +* In `Function.Definitions`, the module 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 two modules have + been removed and their definitions folded into `Function.Definitions`. - - `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` +* In `Function.Definitions` the following definitions 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. -* 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` +* 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)`. -### Change in the definition of `Prime` +### New `Function.Strict` -* 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. +* 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`. -* 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 - ``` +* The contents of `Function.Strict` is now re-exported by `Function`. -### Change to the definition of `Induction.WellFounded.WfRec` (issue #2083) +### Change to the definition of `WfRec` in `Induction.WellFounded` (issue #2083) -* Previously, the following definition was adopted +* Previously, the definition of `WfRec` was: ```agda WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _ WfRec _<_ P x = ∀ y → y < x → P y ``` - with the consequence that all arguments involving about accesibility and - wellfoundedness proofs were polluted by almost-always-inferrable explicit - arguments for the `y` position. The definition has now been changed to - make that argument *implicit*, as + which meant that all arguments involving accessibility and wellfoundedness proofs + were polluted by almost-always-inferrable explicit arguments for the `y` position. + +* The definition has now been changed to make that argument *implicit*, as ```agda WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _ WfRec _<_ P x = ∀ {y} → y < x → P y ``` -### Change in the definition of `_≤″_` (issue #1919) - -* The definition of `_≤″_` in `Data.Nat.Base` was previously: - ```agda - record _≤″_ (m n : ℕ) : Set where - constructor less-than-or-equal - field - {k} : ℕ - proof : m + k ≡ n - ``` - which introduced a spurious additional definition, when this is in fact, modulo - field names and implicit/explicit qualifiers, equivalent to the definition of left- - divisibility, `_∣ˡ_` for the `RawMagma` structure of `_+_`. Since the addition of - raw bundles to `Data.X.Base`, this definition can now be made directly. Knock-on - consequences include the need to retain the old constructor name, now introduced - as a pattern synonym, and introduction of (a function equivalent to) the former - field name/projection function `proof` as `≤″-proof` in `Data.Nat.Properties`. - -* Accordingly, the definition has been changed to: - ```agda - _≤″_ : (m n : ℕ) → Set - _≤″_ = _∣ˡ_ +-rawMagma - - pattern less-than-or-equal {k} prf = k , prf - ``` - -### Renaming of `Reflection` modules +### Reorganisation 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 @@ -681,176 +818,40 @@ Non-backwards compatible changes * 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` + `_>>=_`. - 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. +* 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 about these definitions were spread over + `Relation.Nullary.(Negation/Product/Sum/Implication etc.)`. + +* To fix this all the contents of the latter is now exported by `Relation.Nullary`. + +* In order to achieve this the following backwards compatible changes have been made: + + 1. the definition of `Dec` and `recompute` have been moved to `Relation.Nullary.Decidable.Core` + + 2. the definition of `Reflects` has been moved to `Relation.Nullary.Reflects` + + 3. the definition of `¬_` has been moved to `Relation.Nullary.Negation.Core` + + 4. The modules `Relation.Nullary.(Product/Sum/Implication)` have been deprecated + and their contents moved to `Relation.Nullary.(Negation/Reflects/Decidable)`. + + as well as the following breaking changes: + + 1. `¬?` has been moved from `Relation.Nullary.Negation.Core` to + `Relation.Nullary.Decidable.Core` + + 2. `¬-reflects` has been moved from `Relation.Nullary.Negation.Core` to + `Relation.Nullary.Reflects`. + + 3. `decidable-stable`, `excluded-middle` and `¬-drop-Dec` have been moved + from `Relation.Nullary.Negation` to `Relation.Nullary.Decidable`. + + 4. `fromDec` and `toDec` have been moved from `Data.Sum.Base` to `Data.Sum`. -* 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 - ``` ### (Issue #2096) Introduction of flipped and negated relation symbols to bundles in `Relation.Binary.Bundles` @@ -871,82 +872,22 @@ Non-backwards compatible changes of the old symbol in the record. Therefore, only _declarations_ of `PartialOrder` records will need their field names updating. -### (Issue #1214) Reorganisation of the introduction of negated relation symbols under `Relation.Binary` - -* Previously, negated relation symbols `_≰_` (for `Poset`) and `_≮_` (`TotalOrder`) - were introduced in the corresponding `Relation.Binary.Properties` modules, for re-export. - -* Now they are introduced as definitions in the corresponding `Relation.Binary.Bundles`, - together with, for uniformity's sake, an additional negated symbol `_⋦_` for `Preorder`, - with their (obvious) intended semantics: - ```agda - infix 4 _⋦_ _≰_ _≮_ - Preorder._⋦_ : Rel Carrier _ - StrictPartialOrder._≮_ : Rel Carrier _ - ``` - Additionally, `Poset._≰_` is defined by renaming public export of `Preorder._⋦_` - -* Backwards compatibility has been maintained, with deprecated definitions in the - corresponding `Relation.Binary.Properties` modules, and the corresponding client - client module `import`s being adjusted accordingly. - - -### Standardisation of `insertAt`/`updateAt`/`removeAt` - -* Previously, the names and argument order of index-based insertion, update and removal functions for - various types of lists and vectors were inconsistent. - -* To fix this the names have all been standardised to `insertAt`/`updateAt`/`removeAt`. - -* Correspondingly the following changes have occurred: - -* In `Data.List.Base` the following have been added: - ```agda - insertAt : (xs : List A) → Fin (suc (length xs)) → A → List A - updateAt : (xs : List A) → Fin (length xs) → (A → A) → List A - removeAt : (xs : List A) → Fin (length xs) → List A - ``` - and the following has been deprecated - ``` - _─_ ↦ removeAt - ``` - -* In `Data.Vec.Base`: - ```agda - insert ↦ insertAt - remove ↦ removeAt - - updateAt : Fin n → (A → A) → Vec A n → Vec A n - ↦ - updateAt : Vec A n → Fin n → (A → A) → Vec A n - ``` - -* In `Data.Vec.Functional`: - ```agda - remove : Fin (suc n) → Vector A (suc n) → Vector A n - ↦ - removeAt : Vector A (suc n) → Fin (suc n) → Vector A n - - updateAt : Fin n → (A → A) → Vector A n → Vector A n - ↦ - updateAt : Vector A n → Fin n → (A → A) → Vector A n - ``` - -* The old names (and the names of all proofs about these functions) have been deprecated appropriately. -### Changes to definition of `IsStrictTotalOrder` +### Changes to definition of `IsStrictTotalOrder` in `Relation.Binary.Structures` -* The previous definition of the record `IsStrictTotalOrder` did not build upon `IsStrictPartialOrder` - as would be expected. Instead it omitted several fields like irreflexivity as they were derivable from the - proof of trichotomy. However, this led to problems further up the hierarchy where bundles such as `StrictTotalOrder` - which contained multiple distinct proofs of `IsStrictPartialOrder`. +* The previous definition of the record `IsStrictTotalOrder` did not + build upon `IsStrictPartialOrder` as would be expected. + Instead it omitted several fields like irreflexivity as they were derivable from the + proof of trichotomy. However, this led to problems further up the hierarchy where + bundles such as `StrictTotalOrder` which contained multiple distinct proofs of + `IsStrictPartialOrder`. -* To remedy this the definition of `IsStrictTotalOrder` has been changed to so that it builds upon - `IsStrictPartialOrder` as would be expected. +* To remedy this the definition of `IsStrictTotalOrder` has been changed to so + that it builds upon `IsStrictPartialOrder` as would be expected. -* To aid migration, the old record definition has been moved to `Relation.Binary.Structures.Biased` - which contains the `isStrictTotalOrderᶜ` smart constructor (which is re-exported by `Relation.Binary`) . - Therefore the old code: +* To aid migration, the old record definition has been moved to + `Relation.Binary.Structures.Biased` which contains the `isStrictTotalOrderᶜ` + smart constructor (which is re-exported by `Relation.Binary`) . Therefore the old code: ```agda <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ <-isStrictTotalOrder = record @@ -974,18 +915,20 @@ Non-backwards compatible changes } ``` -### Changes to triple reasoning interface +### Changes to the interface of `Relation.Binary.Reasoning.Triple` -* The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof that the strict - relation is irreflexive. +* 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: +* This allows the addition of 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. + that takes a proof that a value is strictly less than itself and then applies the + principle of explosion to derive anything. -* Specialised versions of this combinator are available in the following derived modules: +* Specialised versions of this combinator are available in the `Reasoning` modules + exported by the following modules: ``` Data.Nat.Properties Data.Nat.Binary.Properties @@ -998,10 +941,14 @@ Non-backwards compatible changes Relation.Binary.Reasoning.PartialOrder ``` -### More modular design of equational reasoning. +### A more modular design for `Relation.Binary.Reasoning` -* Have introduced a new module `Relation.Binary.Reasoning.Syntax` which exports - a range of modules containing pre-existing reasoning combinator syntax. +* Previously, every `Reasoning` module in the library tended to roll it's own set + of syntax for the combinators. This hindered consistency and maintainability. + +* To improve the situation, a new module `Relation.Binary.Reasoning.Syntax` + has been introduced which exports a wide range of sub-modules containing + pre-existing reasoning combinator syntax. * This makes it possible to add new or rename existing reasoning combinators to a pre-existing `Reasoning` module in just a couple of lines @@ -1013,7 +960,7 @@ Non-backwards compatible changes `Relation.Binary.PropositionalEquality.Properties`. It is still exported by `Relation.Binary.PropositionalEquality`. -### Renaming of symmetric combinators for equational reasoning +### Renaming of symmetric combinators in `Reasoning` modules * We've had various complaints about the symmetric version of reasoning combinators that use the syntax `_R˘⟨_⟩_` for some relation `R`, (e.g. `_≡˘⟨_⟩_` and `_≃˘⟨_⟩_`) @@ -1045,59 +992,86 @@ Non-backwards compatible changes `Function.Base` if the `_⟨_⟩_` operator is used within the reasoning proof. The symptom for this is a `Don't know how to parse` error. +### Improvements to `Text.Pretty` and `Text.Regex` + +* 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". + ### 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` +* In `Algebra.Core` the operations `Opₗ` and `Opᵣ` have moved to `Algebra.Module.Core`. -* The first two arguments of `m≡n⇒m-n≡0` (now `i≡j⇒i-j≡0`) in `Data.Integer.Base` - have been made implicit. +* 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. -* The relations `_≤_` `_≥_` `_<_` `_>_` in `Data.Fin.Base` have been +* In `Data.Fin.Base` the relations `_≤_` `_≥_` `_<_` `_>_` 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. +* In `Data.Fin.Properties` the proof `inj⇒≟` 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 and `inj⇒decSetoid`. -* 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-≈`. +* In `Data.Fin.Properties` the proof `pigeonhole` has been strengthened + so that the a proof that `i < j` rather than a mere `i ≢ j` is returned. -* 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 `Data.Fin.Substitution.TermSubst` the fixity of `_/Var_` has been changed + from `infix 8` to `infixl 8`. -* 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. +* In `Data.Integer.DivMod` the previous implementations of + `_divℕ_`, `_div_`, `_modℕ_`, `_mod_` 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. - Also, `⁻¹-homo` in `IsRingHomomorphism` has been renamed to `-‿homo`. +* In `Data.Integer.Properties` the first two arguments of `m≡n⇒m-n≡0` + (now renamed `i≡j⇒i-j≡0`) have been made implicit. -* 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`. +* In `Data.(List|Vec).Relation.Binary.Lex.Strict` the argument `xs` in + `xs≮[]` in introduced in PRs #1648 and #1672 has now been made implicit. -* The functions `split`, `flatten` and `flatten-split` have been removed from - `Data.List.NonEmpty`. In their place `groupSeqs` and `ungroupSeqs` +* In `Data.List.NonEmpty` the functions `split`, `flatten` and `flatten-split` + have been removed from. 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` @@ -1107,138 +1081,89 @@ Non-backwards compatible changes 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)` +* In `Data.Nat.Divisibility` the proof `m/n/o≡m/[n*o]` has been removed. In it's + place a new more general proof `m/n/o≡m/[n*o]` has been added to `Data.Nat.DivMod` + that doesn't require the `n * o ∣ m` pre-condition. + +* In `Data.Product.Relation.Binary.Lex.Strict` the proof of wellfoundedness + of the lexicographic ordering on products, no longer + requires the assumption of symmetry for the first equality relation `_≈₁_`. + +* In `Data.Rational.Base` the constructors `+0` and `+[1+_]` from `Data.Integer.Base` + are no longer re-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: +* In `Data.Rational(.Unnormalised).Properties` the types of the proofs + `pos⇒1/pos`/`1/pos⇒pos` and `neg⇒1/neg`/`1/neg⇒neg` 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: ```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: - ``` + ```agda 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⁻ - ``` +* In `Data.Vec.Base` the definition of `_>>=_` under `Data.Vec.Base` has been + moved to the submodule `CartesianBind` in order to avoid clashing with the + new, alternative definition of `_>>=_`, located in the second new submodule + `DiagonalBind`. -* Renamed in `Data.Vec.Relation.Unary.Linked.Properties` - and `Codata.Guarded.Stream.Relation.Binary.Pointwise`: - ``` - lookup ↦ lookup⁺ - ``` +* In `Data.Vec.Base` the definitions `init` and `last` have been changed from the `initLast` + view-derived implementation to direct recursive definitions. -* Added the following new definitions to `Data.Vec.Relation.Unary.All`: +* In `Data.Vec.Properties` the type of the proof `zipWith-comm` has been generalised from: + ```agda + zipWith-comm : ∀ {f : A → A → B} (comm : ∀ x y → f x y ≡ f y x) (xs ys : Vec A n) → + zipWith f xs ys ≡ zipWith f ys xs ``` - 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) + to + ```agda + zipWith-comm : ∀ {f : A → B → C} {g : B → A → C} (comm : ∀ x y → f x y ≡ g y x) (xs : Vec A n) ys → + zipWith f xs ys ≡ zipWith g ys xs ``` + +* In `Data.Vec.Relation.Unary.All` the functions `lookup` and `tabulate` have + been moved to `Data.Vec.Relation.Unary.All.Properties` and renamed + `lookup⁺` and `lookup⁻` respectively. -* `excluded-middle` in `Relation.Nullary.Decidable.Core` has been renamed to - `¬¬-excluded-middle`. - -* `iterate` and `replicate` in `Data.Vec.Base` and `Data.Vec.Functional` - now take the length of vector, `n`, as an explicit rather than an implicit argument. +* In `Data.Vec.Base` and `Data.Vec.Functional` the functions `iterate` and `replicate` + now take the length of vector, `n`, as an explicit rather than an implicit argument, i.e. + the new types are: ```agda iterate : (A → A) → A → ∀ n → Vec A n replicate : ∀ n → A → Vec A n ``` + +* In `Relation.Binary.Construct.Closure.Symmetric` the operation + `SymClosure` on relations in 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 `Relation.Nullary.Decidable.Core` the name `excluded-middle` has been + renamed to `¬¬-excluded-middle`. -Major improvements ------------------- + +Other 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`. @@ -1256,47 +1181,41 @@ Major improvements * 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 _≈_ _<_` +### Moved raw bundles from `Data.X.Properties` to `Data.X.Base` + +* Raw bundles by design don't contain any proofs so should in theory be able to live + in `Data.X.Base` rather than `Data.X.Bundles`. +* To achieve this while keeping the dependency footprint small, the definitions of + raw bundles (`RawMagma`, `RawMonoid` etc.) have been moved from `Algebra(.Lattice)Bundles` to + a new module `Algebra(.Lattice).Bundles.Raw` which can be imported at much lower cost + from `Data.X.Base`. + +* We have then moved raw bundles defined in `Data.X.Properties` to `Data.X.Base` for + `X` = `Nat`/`Nat.Binary`/`Integer`/`Rational`/`Rational.Unnormalised`. + Deprecated modules ------------------ -### Moving `Relation.Binary.Construct.(Converse/Flip)` +### Moving `Data.Erased` to `Data.Irrelevant` -* The following files have been moved: - ```agda - Relation.Binary.Construct.Converse ↦ Relation.Binary.Construct.Flip.EqAndOrd - Relation.Binary.Construct.Flip ↦ Relation.Binary.Construct.Flip.Ord - ``` +* This fixes the fact we had picked the wrong name originally. The erased modality + corresponds to `@0` whereas the irrelevance one corresponds to `.`. + +### 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.Nat.Properties.Core` + +* The module `Data.Nat.Properties.Core` has been deprecated, and its one lemma moved to `Data.Nat.Base`, renamed as `s≤s⁻¹` + +### 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`. -### Deprecation of old function hierarchy +### Moving `Function.Related` * 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 @@ -1314,56 +1233,15 @@ Deprecated modules Equivalence-kind ↦ EquivalenceKind ``` -### Moving `Algebra.Lattice` files - -* As discussed above the following files have been moved: - ```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 - ``` - -### Moving `*.Catgeorical.*` files +### Moving `Relation.Binary.Construct.(Converse/Flip)` -* As discussed above the following files have been moved: +* The following files have been moved: ```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 + Relation.Binary.Construct.Converse ↦ Relation.Binary.Construct.Flip.EqAndOrd + Relation.Binary.Construct.Flip ↦ Relation.Binary.Construct.Flip.Ord ``` -### Moving `Relation.Binary.Properties.XLattice` files +### Moving `Relation.Binary.Properties.XLattice` * The following files have been moved: ```agda @@ -1377,18 +1255,6 @@ Deprecated modules Relation.Binary.Properties.MeetSemilattice.agda ↦ Relation.Binary.Lattice.Properties.MeetSemilattice.agda ``` -### Deprecation of `Data.Nat.Properties.Core` - -* The module `Data.Nat.Properties.Core` has been deprecated, and its one lemma moved to `Data.Nat.Base`, renamed as `s≤s⁻¹` - -### 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 ---------------- @@ -1438,16 +1304,16 @@ Deprecated names map-map-fusion ↦ map-∘ ``` -* In `Data.Bool.Properties` (Issue #2046): - ``` - push-function-into-if ↦ if-float - ``` - * In `Data.Container.Related`: - ``` + ```agda _∼[_]_ ↦ _≲[_]_ ``` +* In `Data.Bool.Properties` (Issue #2046): + ```agda + push-function-into-if ↦ if-float + ``` + * In `Data.Fin.Base`: two new, hopefully more memorable, names `↑ˡ` `↑ʳ` for the 'left', resp. 'right' injection of a Fin m into a 'larger' type, `Fin (m + n)`, resp. `Fin (n + m)`, with argument order to reflect the @@ -1457,24 +1323,6 @@ Deprecated names raise ↦ _↑ʳ_ ``` - Issue #1726: the relation `_≺_` and its single constructor `_≻toℕ_` - have been deprecated in favour of their extensional equivalent `_<_` - but omitting the inversion principle which pattern matching on `_≻toℕ_` - would achieve; this instead is proxied by the property `Data.Fin.Properties.toℕ<`. - -* In `Data.Fin.Induction`: - ``` - ≺-Rec - ≺-wellFounded - ≺-recBuilder - ≺-rec - ``` - - As with Issue #1726 above: the deprecation of relation `_≺_` means that these definitions - associated with wf-recursion are deprecated in favour of their `_<_` counterparts. - But it's not quite sensible to say that these definitions should be *renamed* to *anything*, - least of all those counterparts... the type confusion would be intolerable. - * In `Data.Fin.Properties`: ``` toℕ-raise ↦ toℕ-↑ʳ @@ -1485,9 +1333,6 @@ Deprecated names eq? ↦ inj⇒≟ ``` - Likewise under issue #1726: the properties `≺⇒<′` and `<′⇒≺` have been deprecated - in favour of their proxy counterparts `<⇒<′` and `<′⇒<`. - * In `Data.Fin.Permutation.Components`: ``` reverse ↦ Data.Fin.Base.opposite @@ -1563,7 +1408,7 @@ Deprecated names ``` * In `Data.List.Base`: - ``` + ```agda _─_ ↦ removeAt ``` @@ -1598,12 +1443,26 @@ Deprecated names * In `Data.List.Relation.Unary.All.Properties`: ```agda + gmap ↦ gmap⁺ + updateAt-id-relative ↦ updateAt-id-local updateAt-compose-relative ↦ updateAt-updateAt-local updateAt-compose ↦ updateAt-updateAt updateAt-cong-relative ↦ updateAt-cong-local ``` +* In `Data.List.Relation.Unary.Any.Properties`: + ```agda + map-with-∈⁺ ↦ mapWith∈⁺ + map-with-∈⁻ ↦ mapWith∈⁻ + map-with-∈↔ ↦ mapWith∈↔ + ``` + +* In `Data.List.Relation.Binary.Subset.Propositional.Properties`: + ```agda + map-with-∈⁺ ↦ mapWith∈⁺ + ``` + * In `Data.List.Zipper.Properties`: ```agda toList-reverse-commute ↦ toList-reverse @@ -1621,26 +1480,10 @@ Deprecated names map-compose ↦ map-∘ map-<∣>-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`: - ``` + ```agda suc[pred[n]]≡n ↦ suc-pred ≤-step ↦ m≤n⇒m≤1+n ≤-stepsˡ ↦ m≤n⇒m≤o+n @@ -1653,14 +1496,14 @@ Deprecated names ``` * In `Data.Rational.Unnormalised.Base`: - ``` + ```agda _≠_ ↦ _≄_ +-rawMonoid ↦ +-0-rawMonoid *-rawMonoid ↦ *-1-rawMonoid ``` * In `Data.Rational.Unnormalised.Properties`: - ``` + ```agda ↥[p/q]≡p ↦ ↥[n/d]≡n ↧[p/q]≡q ↦ ↧[n/d]≡d *-monoˡ-≤-pos ↦ *-monoˡ-≤-nonNeg @@ -1677,13 +1520,13 @@ Deprecated names ``` * In `Data.Rational.Base`: - ``` + ```agda +-rawMonoid ↦ +-0-rawMonoid *-rawMonoid ↦ *-1-rawMonoid ``` * In `Data.Rational.Properties`: - ``` + ```agda *-monoʳ-≤-neg ↦ *-monoʳ-≤-nonPos *-monoˡ-≤-neg ↦ *-monoˡ-≤-nonPos *-monoʳ-≤-pos ↦ *-monoʳ-≤-nonNeg @@ -1697,13 +1540,13 @@ Deprecated names ``` * In `Data.Rational.Unnormalised.Base`: - ``` + ```agda +-rawMonoid ↦ +-0-rawMonoid *-rawMonoid ↦ *-1-rawMonoid ``` * In `Data.Rational.Unnormalised.Properties`: - ``` + ```agda ≤-steps ↦ p≤q⇒p≤r+q ``` @@ -1716,27 +1559,27 @@ Deprecated names ``` * In `Data.Tree.AVL`: - ``` + ```agda _∈?_ ↦ member ``` * In `Data.Tree.AVL.IndexedMap`: - ``` + ```agda _∈?_ ↦ member ``` * In `Data.Tree.AVL.Map`: - ``` + ```agda _∈?_ ↦ member ``` * In `Data.Tree.AVL.Sets`: - ``` + ```agda _∈?_ ↦ member ``` * In `Data.Tree.Binary.Zipper.Properties`: - ``` + ```agda toTree-#nodes-commute ↦ toTree-#nodes toTree-#leaves-commute ↦ toTree-#leaves toTree-map-commute ↦ toTree-map @@ -1751,13 +1594,13 @@ Deprecated names ``` * In `Data.Vec.Base`: - ``` + ```agda remove ↦ removeAt insert ↦ insertAt ``` * In `Data.Vec.Properties`: - ``` + ```agda take-distr-zipWith ↦ take-zipWith take-distr-map ↦ take-map drop-distr-zipWith ↦ drop-zipWith @@ -1787,17 +1630,9 @@ Deprecated names lookup-inject≤-take ↦ lookup-take-inject≤ ``` - and the type of the proof `zipWith-comm` has been generalised from: - ``` - zipWith-comm : ∀ {f : A → A → B} (comm : ∀ x y → f x y ≡ f y x) (xs ys : Vec A n) → zipWith f xs ys ≡ zipWith f ys xs - ``` - to - ``` - zipWith-comm : ∀ {f : A → B → C} {g : B → A → C} (comm : ∀ x y → f x y ≡ g y x) (xs : Vec A n) ys → zipWith f xs ys ≡ zipWith g ys xs - ``` * In `Data.Vec.Functional.Properties`: - ``` + ```agda updateAt-id-relative ↦ updateAt-id-local updateAt-compose-relative ↦ updateAt-updateAt-local updateAt-compose ↦ updateAt-updateAt @@ -1814,7 +1649,7 @@ Deprecated names NB. This last one is complicated by the *addition* of a 'global' property `map-updateAt` * In `Data.Vec.Recursive.Properties`: - ``` + ```agda append-cons-commute ↦ append-cons ``` @@ -1824,12 +1659,12 @@ Deprecated names ``` * In `Function.Base`: - ``` + ```agda case_return_of_ ↦ case_returning_of_ ``` * In `Function.Construct.Composition`: - ``` + ```agda _∘-⟶_ ↦ _⟶-∘_ _∘-↣_ ↦ _↣-∘_ _∘-↠_ ↦ _↠-∘_ @@ -1841,7 +1676,7 @@ Deprecated names ``` * In `Function.Construct.Identity`: - ``` + ```agda id-⟶ ↦ ⟶-id id-↣ ↦ ↣-id id-↠ ↦ ↠-id @@ -1853,7 +1688,7 @@ Deprecated names ``` * In `Function.Construct.Symmetry`: - ``` + ```agda sym-⤖ ↦ ⤖-sym sym-⇔ ↦ ⇔-sym sym-↩ ↦ ↩-sym @@ -1867,60 +1702,142 @@ Deprecated names ``` * In `Foreign.Haskell.Either` and `Foreign.Haskell.Pair`: - ``` + ```agda toForeign ↦ Foreign.Haskell.Coerce.coerce fromForeign ↦ Foreign.Haskell.Coerce.coerce ``` -* In `Relation.Binary.Indexed.Heterogeneous.Bundles.Preorder`: - ``` - _∼_ ↦ _≲_ - ``` - * In `Relation.Binary.Properties.Preorder`: - ``` + ```agda invIsPreorder ↦ converse-isPreorder invPreorder ↦ converse-preorder ``` -* Moved negated relation symbol from `Relation.Binary.Properties.Poset`: - ``` - _≰_ ↦ Relation.Binary.Bundles.Poset._≰_ - ``` - -* Moved negated relation symbol from `Relation.Binary.Properties.TotalOrder`: - ``` - _≮_ ↦ Relation.Binary.Bundles.StrictPartialOrder._≮_ +## Missing fixity declarations added -* In `Relation.Nullary.Decidable.Core`: +* An effort has been made to add sensible fixities for many declarations: ``` - excluded-middle ↦ ¬¬-excluded-middle + infix 4 _≟H_ _≟N_ (Algebra.Solver.Ring) + 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.Guarded.Stream.Relation.Unary.All) + 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) + infix 4 _ℕ<_ _ℕ≤infinity _ℕ≤_ (Codata.Sized.Conat) + infix 6 _ℕ+_ _+ℕ_ (Codata.Sized.Conat) + infixr 8 _⇒_ _⊸_ (Data.Container.Core) + infixr -1 _<$>_ _<*>_ (Data.Container.FreeMonad) + infixl 1 _>>=_ (Data.Container.FreeMonad) + infix 5 _▷_ (Data.Container.Indexed) + infixr 4 _,_ (Data.Container.Relation.Binary.Pointwise) + infix 4 _≈_ (Data.Float.Base) + 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) + infixl 7 _⊓′_ (Data.Nat.Base) + infixl 6 _⊔′_ (Data.Nat.Base) + infixr 8 _^_ (Data.Nat.Base) + infix 4 _!≢0 _!*_!≢0 (Data.Nat.Properties) + infixl 6.5 _P′_ _P_ _C′_ _C_ (Data.Nat.Combinatorics.Base) + infix 8 _⁻¹ (Data.Parity.Base) + 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) + infix 4 _≃?_ (Data.Rational.Unnormalised.Properties) + infixr 4 _,_ (Data.Refinement) + 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) + infixr 4 _,_ (Data.Tree.AVL.Value) + infix 4 _≈ₖᵥ_ (Data.Tree.AVL.Map.Membership.Propositional) + infixr 5 _`∷_ (Data.Vec.Reflection) + infixr 5 _∷=_ (Data.Vec.Membership.Setoid) + infix -1 _$ⁿ_ (Data.Vec.N-ary) + infix 4 _≋_ (Data.Vec.Functional.Relation.Binary.Equality.Setoid) + infixl 1 _>>=-cong_ _≡->>=-cong_ (Effect.Monad.Partiality) + infixl 1 _?>=′_ (Effect.Monad.Predicate) + infixl 1 _>>=-cong_ _>>=-congP_ (Effect.Monad.Partiality.All) + infixr 5 _∷_ (Foreign.Haskell.List.NonEmpty) + infixr 4 _,_ (Foreign.Haskell.Pair) + 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 4 _≈_ (Function.Metric.Rational.Bundles) + infix 3 _←_ _↢_ (Function.Related) + infix 4 _<_ (Induction.WellFounded) + infixl 6 _ℕ+_ (Level.Literals) + infixr 4 _,_ (Reflection.AnnotatedAST) + 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 4 _,_ (Reflection.AST.Traversal) + infix 4 _∈FV_ (Reflection.AST.DeBruijn) + infixr 9 _;_ (Relation.Binary.Construct.Composition) + infixl 6 _+²_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) + infixr -1 _atₛ_ (Relation.Binary.Indexed.Heterogeneous.Construct.At) + infixr -1 _atₛ_ (Relation.Binary.Indexed.Homogeneous.Construct.At) + infix 4 _∈_ _∉_ (Relation.Unary.Indexed) + 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 6 _∪_ (Relation.Binary.Construct.Union) + infixl 6 _+ℤ_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) + infix 4 _≉_ _≈ᵢ_ _≤ᵢ_ (Relation.Binary.Indexed.Homogeneous.Bundles) + infixr 9 _⍮_ (Relation.Unary.PredicateTransformer) + infix 8 ∼_ (Relation.Unary.PredicateTransformer) + infix 2 _×?_ _⊙?_ (Relation.Unary.Properties) + infix 10 _~? (Relation.Unary.Properties) + infixr 1 _⊎?_ (Relation.Unary.Properties) + infixr 7 _∩?_ (Relation.Unary.Properties) + infixr 6 _∪?_ (Relation.Unary.Properties) + infixr 5 _∷ᴹ_ _∷⁻¹ᴹ_ (Text.Regex.Search) + infixl 6 _`⊜_ (Tactic.RingSolver) + infix 8 ⊝_ (Tactic.RingSolver.Core.Expression) + infix 4 _∈ᴿ?_ _∉ᴿ?_ _∈?ε _∈?[_] _∈?[^_] (Text.Regex.Properties) + infix 4 _∈?_ _∉?_ (Text.Regex.Derivative.Brzozowski) + infix 4 _∈_ _∉_ _∈?_ _∉?_ (Text.Regex.String.Unsafe) ``` -### Renamed Data.Erased to Data.Irrelevant - -* This fixes the fact we had picked the wrong name originally. The erased modality - corresponds to @0 whereas the irrelevance one corresponds to `.`. - - New modules ----------- -* Algebraic structures when freely adding an identity element: - ``` - Algebra.Construct.Add.Identity - ``` - -* Operations for module-like algebraic structures: - ``` - Algebra.Module.Core - ``` - -* Definitions for module-like algebraic structures with left- and right- multiplication over the same scalars: - ``` - Algebra.Module.Definitions.Bi.Simultaneous - ``` - * Constructive algebraic structures with apartness relations: ``` Algebra.Apartness @@ -1935,11 +1852,20 @@ New modules Algebra.Construct.Flip.Op ``` -* Morphisms between module-like algebraic structures: +* Algebraic structures when freely adding an identity element: + ``` + Algebra.Construct.Add.Identity + ``` + +* Definitions for algebraic modules: ``` + Algebra.Module + Algebra.Module.Core + Algebra.Module.Definitions.Bi.Simultaneous Algebra.Module.Morphism.Construct.Composition Algebra.Module.Morphism.Construct.Identity Algebra.Module.Morphism.Definitions + Algebra.Module.Morphism.ModuleHomomorphism Algebra.Module.Morphism.Structures Algebra.Module.Properties ``` @@ -1950,11 +1876,36 @@ New modules Algebra.Morphism.Construct.Identity ``` +* Properties of various new algebraic structures: + ``` + Algebra.Properties.MoufangLoop + Algebra.Properties.Quasigroup + Algebra.Properties.MiddleBolLoop + Algebra.Properties.Loop + Algebra.Properties.KleeneAlgebra + ``` + +* Properties of rings without a unit + ``` + Algebra.Properties.RingWithoutOne` + ``` + +* Proof of the Binomial Theorem for semirings + ``` + Algebra.Properties.Semiring.Binomial + Algebra.Properties.CommutativeSemiring.Binomial + ``` + * 'Optimised' tail-recursive exponentiation properties: ``` Algebra.Properties.Semiring.Exp.TailRecursiveOptimised ``` +* An implementation of M-types with `--guardedness` flag: + ``` + Codata.Guarded.M + ``` + * A definition of infinite streams using coinductive records: ``` Codata.Guarded.Stream @@ -2015,6 +1966,17 @@ New modules Data.List.NonEmpty.Relation.Unary.All ``` +* Some n-ary functions manipulating lists + ``` + Data.List.Nary.NonDependent + ``` + +* Added Logarithm base 2 on natural numbers: + ``` + Data.Nat.Logarithm.Core + Data.Nat.Logarithm + ``` + * Show module for unnormalised rationals: ``` Data.Rational.Unnormalised.Show @@ -2028,7 +1990,23 @@ New modules Data.Tree.AVL.Sets.Membership.Properties ``` -* Properties of bijections: +* Port of `Linked` to `Vec`: + ``` + Data.Vec.Relation.Unary.Linked + Data.Vec.Relation.Unary.Linked.Properties + ``` + +* Combinators for propositional equational reasoning on vectors with different indices + ``` + Data.Vec.Relation.Binary.Equality.Cast + ``` + +* Relations on indexed sets + ``` + Function.Indexed.Bundles + ``` + +* Properties of various types of functions: ``` Function.Consequences Function.Properties.Bijection @@ -2037,6 +2015,11 @@ New modules Function.Construct.Constant ``` +* New interface for `NonEmpty` Haskell lists: + ``` + Foreign.Haskell.List.NonEmpty + ``` + * The 'no infinite descent' principle for (accessible elements of) well-founded relations: ``` Induction.NoInfiniteDescent @@ -2051,6 +2034,12 @@ New modules ``` All contents is re-exported by `Relation.Binary.Lattice` as before. + +* Added relational reasoning over apartness relations: + ``` + Relation.Binary.Reasoning.Base.Apartness` + ``` + * Algebraic properties of `_∩_` and `_∪_` for predicates ``` Relation.Unary.Algebra @@ -2077,11 +2066,6 @@ New modules Reflection.AST.AlphaEquality ``` -* `cong!` tactic for deriving arguments to `cong` - ``` - Tactic.Cong - ``` - * Various system types and primitives: ``` System.Clock.Primitive @@ -2095,161 +2079,94 @@ New modules System.Process ``` -* A golden testing library with test pools, an options parser, a runner: - ``` - Test.Golden - ``` - -* New interfaces for using Haskell datatypes: - ``` - Foreign.Haskell.List.NonEmpty - ``` -* Added new module `Algebra.Properties.RingWithoutOne`: - ``` - -‿distribˡ-* : ∀ x y → - (x * y) ≈ - x * y - -‿distribʳ-* : ∀ x y → - (x * y) ≈ x * - y - ``` -* An implementation of M-types with `--guardedness` flag: - ``` - Codata.Guarded.M - ``` - -* Port of `Linked` to `Vec`: - ``` - Data.Vec.Relation.Unary.Linked - Data.Vec.Relation.Unary.Linked.Properties - ``` -* Added Logarithm base 2 on Natural Numbers: - ``` - Data.Nat.Logarithm.Core - Data.Nat.Logarithm - ``` - -* Proofs of some axioms of linearity: - ``` - Algebra.Module.Morphism.ModuleHomomorphism - Algebra.Module.Properties - ``` -* Properties of MoufangLoop - ``` - Algebra.Properties.MoufangLoop - ``` - -* Properties of Quasigroup - ``` - Algebra.Properties.Quasigroup - ``` - -* Properties of MiddleBolLoop - ``` - Algebra.Properties.MiddleBolLoop - ``` - -* Properties of Loop - ``` - Algebra.Properties.Loop - ``` - -* Properties of (Commutative)Semiring: the Binomial Theorem - ``` - Algebra.Properties.CommutativeSemiring.Binomial - Algebra.Properties.Semiring.Binomial - ``` - -* Some n-ary functions manipulating lists - ``` - Data.List.Nary.NonDependent - ``` - -* Properties of KleeneAlgebra - ``` - Algebra.Properties.KleeneAlgebra - ``` - -* Relations on indexed sets +* A new `cong!` tactic for automatically deriving arguments to `cong` ``` - Function.Indexed.Bundles + Tactic.Cong ``` -* Combinators for propositional equational reasoning on vectors with different indices +* A golden testing library with test pools, an options parser, a runner: ``` - Data.Vec.Relation.Binary.Equality.Cast + Test.Golden ``` Additions to existing modules ----------------------------- -* Added new proof to `Data.Maybe.Properties` - ```agda - <∣>-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 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 + record Quasigroup 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 Quasiring c ℓ : Set (suc (c ⊔ ℓ)) + record Nearring c ℓ : Set (suc (c ⊔ ℓ)) + 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 ⊔ ℓ)) + ``` + +* Added new definitions to `Algebra.Bundles.Raw`: + ```agda + record RawLoop c ℓ : Set (suc (c ⊔ ℓ)) + record RawQuasiGroup c ℓ : Set (suc (c ⊔ ℓ)) + record RawRingWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) + ``` + +* Added new definitions to `Algebra.Structures`: ```agda - ∨-commutativeSemigroup : CommutativeSemigroup c ℓ - ∧-commutativeSemigroup : CommutativeSemigroup c ℓ + 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 their corresponding algebraic subbundles. -* Added new proofs to `Algebra.Consequences.Base`: +* Added new proof to `Algebra.Consequences.Base`: ```agda - reflexive+selfInverse⇒involutive : Reflexive _≈_ → - SelfInverse _≈_ f → - Involutive _≈_ f + 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 _•_ + 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 _•_ + 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 @@ -2272,63 +2189,60 @@ Additions to existing modules * Added new functions to `Algebra.Construct.DirectProduct`: ```agda - rawSemiring : RawSemiring a ℓ₁ → RawSemiring b ℓ₂ → RawSemiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - rawRing : RawRing a ℓ₁ → RawRing b ℓ₂ → RawRing (a ⊔ b) (ℓ₁ ⊔ ℓ₂) + 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) (ℓ₁ ⊔ ℓ₂) - ``` + 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 _ + _*_ MiddleFourExchange _+_ = ∀ w x y z → ((w + x) * (y + z)) ≈ ((w + y) * (x + z)) - SelfInverse : Op₁ A → Set _ + SelfInverse f = ∀ {x y} → f x ≈ y → f y ≈ x - 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 _ + LeftDividesˡ _∙_ _\\_ = ∀ x y → (x ∙ (x \\ y)) ≈ y + LeftDividesʳ _∙_ _\\_ = ∀ x y → (x \\ (x ∙ y)) ≈ y + RightDividesˡ _∙_ _//_ = ∀ x y → ((y // x) ∙ x) ≈ y + RightDividesʳ _∙_ _//_ = ∀ x y → ((y ∙ x) // x) ≈ y + LeftDivides ∙ \\ = (LeftDividesˡ ∙ \\) × (LeftDividesʳ ∙ \\) + RightDivides ∙ // = (RightDividesˡ ∙ //) × (RightDividesʳ ∙ //) LeftInvertible e _∙_ x = ∃[ x⁻¹ ] (x⁻¹ ∙ x) ≈ e RightInvertible e _∙_ x = ∃[ x⁻¹ ] (x ∙ x⁻¹) ≈ e @@ -2355,24 +2269,64 @@ Additions to existing modules * Added new functions to `Algebra.Definitions.RawSemiring`: ```agda _^[_]*_ : A → ℕ → A → A - _^ᵗ_ : A → ℕ → A + _^ᵗ_ : A → ℕ → A + ``` + +* In `Algebra.Bundles.Lattice` the existing record `Lattice` now provides + ```agda + ∨-commutativeSemigroup : CommutativeSemigroup c ℓ + ∧-commutativeSemigroup : CommutativeSemigroup c ℓ + ``` + and their corresponding algebraic sub-bundles. + +* In `Algebra.Lattice.Structures` the record `IsLattice` now provides + ``` + ∨-isCommutativeSemigroup : IsCommutativeSemigroup ∨ + ∧-isCommutativeSemigroup : IsCommutativeSemigroup ∧ ``` + and their corresponding algebraic substructures. -* `Algebra.Properties.Magma.Divisibility` now re-exports operations +* The module `Algebra.Properties.Magma.Divisibility` now re-exports operations `_∣ˡ_`, `_∤ˡ_`, `_∣ʳ_`, `_∤ʳ_` from `Algebra.Definitions.Magma`. +* 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 to `Algebra.Properties.CommutativeSemigroup`: ```agda + xy∙xx≈x∙yxx : (x ∙ y) ∙ (x ∙ x) ≈ x ∙ (y ∙ (x ∙ x)) + 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) + middleSemimedial : (x ∙ y) ∙ (z ∙ x) ≈ (x ∙ z) ∙ (y ∙ x) semimedial : Semimedial _∙_ ``` + +* 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 proof to `Algebra.Properties.Monoid.Mult`: ```agda - ×-congˡ : ∀ {x} → (_× x) Preserves _≡_ ⟶ _≈_ + ×-congˡ : (_× x) Preserves _≡_ ⟶ _≈_ ``` * Added new proof to `Algebra.Properties.Monoid.Sum`: @@ -2382,16 +2336,16 @@ Additions to existing modules * Added new proofs to `Algebra.Properties.Semigroup`: ```agda - leftAlternative : LeftAlternative _∙_ + leftAlternative : LeftAlternative _∙_ rightAlternative : RightAlternative _∙_ - alternative : Alternative _∙_ - flexible : Flexible _∙_ + 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 + 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`: @@ -2403,56 +2357,25 @@ Additions to existing modules * 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) + -1*x≈-x : - 1# * x ≈ - x + x+x≈x⇒x≈0 : x + x ≈ x → x ≈ 0# + x[y-z]≈xy-xz : x * (y - z) ≈ x * y - x * z + [y-z]x≈yx-zx : (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 +* Added new functions in `Codata.Guarded.Stream`: ``` - ∨-isCommutativeSemigroup : IsCommutativeSemigroup ∨ - ∧-isCommutativeSemigroup : IsCommutativeSemigroup ∧ + transpose : List (Stream A) → Stream (List A) + transpose⁺ : List⁺ (Stream A) → Stream (List⁺ A) + concat : Stream (List⁺ A) → Stream A ``` - 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 `Codata.Guarded.Stream.Properties`: + ``` + cong-concat : ass ≈ bss → concat ass ≈ concat bss + map-concat : map f (concat ass) ≈ concat (map (List⁺.map f) ass) + lookup-transpose : lookup n (transpose ass) ≡ List.map (lookup n) ass + lookup-transpose⁺ : lookup n (transpose⁺ ass) ≡ List⁺.map (lookup n) ass ``` * Added new proofs in `Data.Bool.Properties`: @@ -2483,6 +2406,25 @@ Additions to existing modules xor-annihilates-not : (not x) xor (not y) ≡ x xor y ``` +* Exposed container combinator conversion functions from `Data.Container.Combinator.Properties` + separately 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 → ⟦ const A ⟧ B + from-const : ⟦ const A ⟧ B → A + to-∘ : ⟦ C₁ ⟧ (⟦ C₂ ⟧ A) → ⟦ C₁ ∘ C₂ ⟧ A + from-∘ : ⟦ C₁ ∘ C₂ ⟧ A → ⟦ C₁ ⟧ (⟦ C₂ ⟧ A) + to-× : ⟦ C₁ ⟧ A P.× ⟦ C₂ ⟧ A → ⟦ C₁ × C₂ ⟧ A + from-× : ⟦ C₁ × C₂ ⟧ A → ⟦ C₁ ⟧ A P.× ⟦ C₂ ⟧ A + to-Π : (∀ i → ⟦ Cᵢ i ⟧ A) → ⟦ Π I Cᵢ ⟧ A + from-Π : ⟦ Π I Cᵢ ⟧ A → ∀ i → ⟦ Cᵢ i ⟧ A + to-⊎ : ⟦ C₁ ⟧ A S.⊎ ⟦ C₂ ⟧ A → ⟦ C₁ ⊎ C₂ ⟧ A + from-⊎ : ⟦ C₁ ⊎ C₂ ⟧ A → ⟦ C₁ ⟧ A S.⊎ ⟦ C₂ ⟧ A + to-Σ : (∃ λ i → ⟦ C i ⟧ A) → ⟦ Σ I C ⟧ A + from-Σ : ⟦ Σ I C ⟧ A → ∃ λ i → ⟦ C i ⟧ A + ``` + * Added new functions in `Data.Fin.Base`: ``` finToFun : Fin (m ^ n) → (Fin n → Fin m) @@ -2492,11 +2434,11 @@ Additions to existing modules ``` * 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 _⊏_) + spo-wellFounded : IsStrictPartialOrder _≈_ _⊏_ → WellFounded _⊏_ + spo-noetherian : IsStrictPartialOrder _≈_ _⊏_ → WellFounded (flip _⊏_) + + <-weakInduction-startingFrom : P i → (∀ j → P (inject₁ j) → P (suc j)) → ∀ {j} → j ≥ i → P j ``` * Added new definitions and proofs in `Data.Fin.Permutation`: @@ -2507,15 +2449,10 @@ Additions to existing modules 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 ↔ ⊤ + 2↔Bool : Fin 2 ↔ Bool 0≢1+n : zero ≢ suc i @@ -2560,31 +2497,24 @@ Additions to existing modules 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 + i≤inject₁[j]⇒i≤1+j : i ≤ inject₁ j → i ≤ suc j ``` * 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`: + map-var≡ : (∀ x → lookup ρ₁ x ≡ f x) → (∀ x → lookup ρ₂ x ≡ T.var (f x)) → map var ρ₁ ≡ ρ₂ + wk≡wk : map var VarSubst.wk ≡ T.wk {n = n} + id≡id : map var VarSubst.id ≡ T.id {n = n} + sub≡sub : map var (VarSubst.sub x) ≡ T.sub (T.var x) + ↑≡↑ : map var (ρ VarSubst.↑) ≡ map T.var ρ T.↑ + /Var≡/ : t /Var ρ ≡ t T./ map T.var ρ + + sub-renaming-commutes : t /Var VarSubst.sub x T./ ρ ≡ t T./ ρ T.↑ T./ T.sub (lookup ρ x) + sub-commutes-with-renaming : t T./ T.sub t′ /Var ρ ≡ t /Var ρ VarSubst.↑ T./ T.sub (t′ /Var ρ) ``` + +* Added new functions and definitions in `Data.Integer.Base`: + ```agda _^_ : ℤ → ℕ → ℤ +-0-rawGroup : Rawgroup 0ℓ 0ℓ @@ -2659,11 +2589,32 @@ Additions to existing modules ∈-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) → (ix : x ∈ xs) → ∈-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix ``` +* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`: + ```agda + foldr-commMonoid : xs ↭ ys → foldr _∙_ ε xs ≈ foldr _∙_ ε ys + ``` + +* Added new function to `Data.List.Relation.Binary.Permutation.Propositional.Properties` + ```agda + ↭-reverse : reverse xs ↭ xs + ``` + +* Added new proofs to `Data.List.Relation.Binary.Sublist.Setoid.Properties`: + ``` + ⊆-mergeˡ : xs ⊆ merge _≤?_ xs ys + ⊆-mergeʳ : ys ⊆ merge _≤?_ xs ys + ``` + * Added new functions in `Data.List.Relation.Unary.All`: ``` decide : Π[ P ∪ Q ] → Π[ All P ∪ Any Q ] ``` +* Added new proof to `Data.List.Relation.Unary.All.Properties`: + ```agda + gmap⁻ : Q ∘ f ⋐ P → All Q ∘ map f ⋐ All P + ``` + * Added new functions in `Data.List.Fresh.Relation.Unary.All`: ``` decide : Π[ P ∪ Q ] → Π[ All {R = R} P ∪ Any Q ] @@ -2677,15 +2628,22 @@ Additions to existing modules * 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) + 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₂ + + ∈-++⁺∘++⁻ : (p : v ∈ xs ++ ys) → [ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ (∈-++⁻ xs p) ≡ p + ∈-++⁻∘++⁺ : (p : v ∈ xs ⊎ v ∈ ys) → ∈-++⁻ xs ([ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ p) ≡ p + ∈-++↔ : (v ∈ xs ⊎ v ∈ ys) ↔ v ∈ xs ++ ys + ∈-++-comm : v ∈ xs ++ ys → v ∈ ys ++ xs + ∈-++-comm∘++-comm : (p : v ∈ xs ++ ys) → ∈-++-comm ys xs (∈-++-comm xs ys p) ≡ p + ∈-++↔++ : v ∈ xs ++ ys ↔ v ∈ ys ++ xs ``` * Add new proofs in `Data.List.Properties`: ```agda ∈⇒∣product : n ∈ ns → n ∣ product ns - ∷ʳ-++ : xs ∷ʳ a ++ ys ≡ xs ++ a ∷ ys + ∷ʳ-++ : xs ∷ʳ a ++ ys ≡ xs ++ a ∷ ys concatMap-cong : f ≗ g → concatMap f ≗ concatMap g concatMap-pure : concatMap [_] ≗ id @@ -2716,14 +2674,46 @@ Additions to existing modules map-replicate : map f (replicate n x) ≡ replicate n (f x) zipWith-replicate : zipWith _⊕_ (replicate n x) (replicate n y) ≡ replicate n (x ⊕ y) - length-iterate : length (iterate f x n) ≡ n - iterate-id : iterate id x n ≡ replicate n x - lookup-iterate : lookup (iterate f x n) (cast (sym (length-iterate f x n)) i) ≡ ℕ.iterate f x (toℕ i) + length-iterate : length (iterate f x n) ≡ n + iterate-id : iterate id x n ≡ replicate n x + lookup-iterate : lookup (iterate f x n) (cast (sym (length-iterate f x n)) i) ≡ ℕ.iterate f x (toℕ i) + + length-insertAt : length (insertAt xs i v) ≡ suc (length xs) + length-removeAt′ : length xs ≡ suc (length (removeAt xs k)) + removeAt-insertAt : removeAt (insertAt xs i v) ((cast (sym (length-insertAt xs i v)) i)) ≡ xs + insertAt-removeAt : insertAt (removeAt xs i) (cast (sym (lengthAt-removeAt xs i)) i) (lookup xs i) ≡ xs + + 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 + ``` + +* In `Data.List.NonEmpty.Base`: + ```agda + drop+ : ℕ → List⁺ A → List⁺ A + ``` + +* 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 + ``` - length-insertAt : length (insertAt xs i v) ≡ suc (length xs) - length-removeAt′ : length xs ≡ suc (length (removeAt xs k)) - removeAt-insertAt : removeAt (insertAt xs i v) ((cast (sym (length-insertAt xs i v)) i)) ≡ xs - insertAt-removeAt : insertAt (removeAt xs i) (cast (sym (lengthAt-removeAt xs i)) i) (lookup xs i) ≡ xs +* Added new proof to `Data.Maybe.Properties` + ```agda + <∣>-idem : Idempotent _<∣>_ ``` * Added new patterns and definitions to `Data.Nat.Base`: @@ -2740,8 +2730,8 @@ Additions to existing modules pattern ≤″-offset k = less-than-or-equal {k} refl pattern <″-offset k = ≤″-offset k - s≤″s⁻¹ : ∀ {m n} → suc m ≤″ suc n → m ≤″ n - s<″s⁻¹ : ∀ {m n} → suc m <″ suc n → m <″ n + s≤″s⁻¹ : suc m ≤″ suc n → m ≤″ n + s<″s⁻¹ : suc m <″ suc n → m <″ n _⊔′_ : ℕ → ℕ → ℕ _⊓′_ : ℕ → ℕ → ℕ @@ -2795,7 +2785,7 @@ Additions to existing modules m*n≢0 : .{{_ : NonZero m}} .{{_ : NonZero n}} → NonZero (m * n) m≤n⇒n∸m≤n : m ≤ n → n ∸ m ≤ n - s0 : ∀ m .{{_ : NonZero m}} n → m ^ n > 0 + m^n>0 : .{{_ : 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 _<_ ⟶ _<_ + ^-monoˡ-≤ : (_^ n) Preserves _≤_ ⟶ _≤_ + ^-monoʳ-≤ : .{{_ : NonZero m}} → (m ^_) Preserves _≤_ ⟶ _≤_ + ^-monoˡ-< : .{{_ : NonZero n}} → (_^ n) Preserves _<_ ⟶ _<_ + ^-monoʳ-< : 1 < m → (m ^_) Preserves _<_ ⟶ _<_ n≡⌊n+n/2⌋ : n ≡ ⌊ n + n /2⌋ n≡⌈n+n/2⌉ : n ≡ ⌈ n + n /2⌉ @@ -2840,10 +2830,7 @@ Additions to existing modules ⊔≡⊔′ : m ⊔ n ≡ m ⊔′ n ⊓≡⊓′ : m ⊓ n ≡ m ⊓′ n ∣-∣≡∣-∣′ : ∣ m - n ∣ ≡ ∣ m - n ∣′ - ``` - -* Re-exported additional functions from `Data.Nat`: - ```agda + nonZero? : Decidable NonZero eq? : A ↣ ℕ → DecidableEquality A ≤-<-connex : Connex _≤_ _<_ @@ -2859,7 +2846,7 @@ Additions to existing modules ```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 ! + 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 @@ -2914,6 +2901,12 @@ Additions to existing modules pattern `suc x = con (quote ℕ.suc) (x ⟨∷⟩ []) ``` +* Added new functions and proofs in `Data.Nat.GeneralisedArithmetic`: + ```agda + iterate : (A → A) → A → ℕ → A + iterate-is-fold : fold z s m ≡ iterate s z m + ``` + * Added new proofs in `Data.Parity.Properties`: ```agda suc-homo-⁻¹ : (parity (suc n)) ⁻¹ ≡ parity n @@ -2933,8 +2926,8 @@ Additions to existing modules * Added new definitions and proofs in `Data.Rational.Properties`: ```agda - ↥ᵘ-toℚᵘ : ↥ᵘ (toℚᵘ p) ≡ ↥ p - ↧ᵘ-toℚᵘ : ↧ᵘ (toℚᵘ p) ≡ ↧ p + ↥ᵘ-toℚᵘ : ↥ᵘ (toℚᵘ p) ≡ ↥ p + ↧ᵘ-toℚᵘ : ↧ᵘ (toℚᵘ p) ≡ ↧ p ↥p≡↥q≡0⇒p≡q : ↥ p ≡ 0ℤ → ↥ q ≡ 0ℤ → p ≡ q _≥?_ : Decidable _≥_ @@ -2950,6 +2943,10 @@ Additions to existing modules pos⇒nonZero : .{{Positive p}} → NonZero p neg⇒nonZero : .{{Negative p}} → NonZero p nonZero⇒1/nonZero : .{{_ : NonZero p}} → NonZero (1/ p) + + <-dense : Dense _<_ + <-isDenseLinearOrder : IsDenseLinearOrder _≡_ _<_ + <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ ``` * Added new rounding functions in `Data.Rational.Unnormalised.Base`: @@ -2960,8 +2957,8 @@ Additions to existing modules * Added new definitions in `Data.Rational.Unnormalised.Properties`: ```agda - ↥p≡0⇒p≃0 : ↥ p ≡ 0ℤ → p ≃ 0ℚᵘ - p≃0⇒↥p≡0 : p ≃ 0ℚᵘ → ↥ p ≡ 0ℤ + ↥p≡0⇒p≃0 : ↥ p ≡ 0ℤ → p ≃ 0ℚᵘ + p≃0⇒↥p≡0 : p ≃ 0ℚᵘ → ↥ p ≡ 0ℤ ↥p≡↥q≡0⇒p≃q : ↥ p ≡ 0ℤ → ↥ q ≡ 0ℤ → p ≃ q +-*-rawNearSemiring : RawNearSemiring 0ℓ 0ℓ @@ -2987,12 +2984,29 @@ Additions to existing modules 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) + + 0≄1 : 0ℚᵘ ≄ 1ℚᵘ + ≃-≄-irreflexive : Irreflexive _≃_ _≄_ + ≄-symmetric : Symmetric _≄_ + ≄-cotransitive : Cotransitive _≄_ + ≄⇒invertible : p ≄ q → Invertible _≃_ 1ℚᵘ _*_ (p - q) + + <-dense : Dense _<_ + + <-isDenseLinearOrder : IsDenseLinearOrder _≃_ _<_ + +-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ + +-*-isHeytingField : IsHeytingField _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ + + <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ + +-*-heytingCommutativeRing : HeytingCommutativeRing 0ℓ 0ℓ 0ℓ + +-*-heytingField : HeytingField 0ℓ 0ℓ 0ℓ + + module ≃-Reasoning = SetoidReasoning ≃-setoid ``` * 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 + 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`: @@ -3000,8 +3014,7 @@ Additions to existing modules 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: +* Added new definitions to `Data.Product.Properties`: ``` Σ-≡,≡→≡ : (∃ λ (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₂) @@ -3011,12 +3024,9 @@ Additions to existing modules * 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 _<ₗₑₓ_ + ×-respectsʳ : Transitive _≈₁_ → _<₁_ Respectsʳ _≈₁_ → _<₂_ Respectsʳ _≈₂_ → _<ₗₑₓ_ Respectsʳ _≋_ + ×-respectsˡ : Symmetric _≈₁_ → Transitive _≈₁_ → _<₁_ Respectsˡ _≈₁_ → _<₂_ Respectsˡ _≈₂_ → _<ₗₑₓ_ Respectsˡ _≋_ + ×-wellFounded' : Transitive _≈₁_ → _<₁_ Respectsʳ _≈₁_ → WellFounded _<₁_ → WellFounded _<₂_ → WellFounded _<ₗₑₓ_ ``` * Added new definitions to `Data.Sign.Base`: @@ -3047,21 +3057,19 @@ Additions to existing modules ``` * Added new proofs in `Data.String.Properties`: - ``` + ```agda ≤-isDecTotalOrder-≈ : IsDecTotalOrder _≈_ _≤_ - ≤-decTotalOrder-≈ : DecTotalOrder _ _ _ + ≤-decTotalOrder-≈ : DecTotalOrder _ _ _ ``` * Added new definitions in `Data.Sum.Properties`: - ``` + ```agda 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 + ```agda + map-assocˡ : map (map f g) h ∘ assocˡ ≗ assocˡ ∘ map f (map g h) + map-assocʳ : map f (map g h) ∘ assocʳ ≗ assocʳ ∘ map (map f g) h ``` * Made `Map` public in `Data.Tree.AVL.IndexedMap` @@ -3071,8 +3079,8 @@ Additions to existing modules 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) + FoldrOp A B = A → B n → B (suc n) + FoldlOp A B = B n → A → B (suc n) foldr′ : (A → B → B) → B → Vec A n → B foldl′ : (B → A → B) → B → Vec A n → B @@ -3121,7 +3129,7 @@ Additions to existing modules []≔-++-↑ʳ : (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) → + foldl-universal : (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 @@ -3207,6 +3215,11 @@ Additions to existing modules index-∈-fromList⁺ : Any.index (∈-fromList⁺ v∈xs) ≡ indexₗ v∈xs ``` +* Added new isomorphisms to `Data.Unit.Polymorphic.Properties`: + ```agda + ⊤↔⊤* : ⊤ ↔ ⊤* + ``` + * 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) @@ -3214,13 +3227,21 @@ Additions to existing modules * 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}) + xs≮[] : ¬ xs < [] + + <-respectsˡ : IsPartialEquivalence _≈_ → _≺_ Respectsˡ _≈_ → _<_ Respectsˡ _≋_ + <-respectsʳ : IsPartialEquivalence _≈_ → _≺_ Respectsʳ _≈_ → _<_ _Respectsʳ _≋_ + <-wellFounded : Transitive _≈_ → _≺_ Respectsʳ _≈_ → WellFounded _≺_ → WellFounded _<_ + ``` + +* 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 functions in `Data.Vec.Relation.Unary.Any`: @@ -3231,12 +3252,28 @@ Additions to existing modules * Added new functions in `Data.Vec.Relation.Unary.All`: ``` decide : Π[ P ∪ Q ] → Π[ All P ∪ Any Q ] + + 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) ``` * Added vector associativity proof to `Data.Vec.Relation.Binary.Equality.Setoid`: ``` ++-assoc : (xs ++ ys) ++ zs ≋ xs ++ (ys ++ zs) ``` + +* 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↔ : A ↔ B → A ^ n ↔ B ^ n + Fin[m^n]↔Fin[m]^n : Fin (m ^ n) ↔ Fin m Vec.^ n + ``` * Added new functions in `Effect.Monad.State`: ``` @@ -3245,17 +3282,29 @@ Additions to existing modules execState : State s a → s → s ``` +* Added a non-dependent version of `flip` in `Function.Base`: + ```agda + flip′ : (A → B → C) → (B → A → C) + ``` + * Added new proofs and definitions in `Function.Bundles`: ```agda LeftInverse.isSplitSurjection : LeftInverse → IsSplitSurjection to LeftInverse.surjection : LeftInverse → Surjection SplitSurjection = LeftInverse + _⟨$⟩_ = Func.to ``` -* Added new application operator synonym to `Function.Bundles`: +* Added new proofs to `Function.Properties.Inverse`: ```agda - _⟨$⟩_ : Func From To → Carrier From → Carrier To - _⟨$⟩_ = Func.to + ↔-refl : Reflexive _↔_ + ↔-sym : Symmetric _↔_ + ↔-trans : Transitive _↔_ + + ↔⇒↣ : A ↔ B → A ↣ B + ↔-fun : A ↔ B → C ↔ D → (A → C) ↔ (B → D) + + Inverse⇒Injection : Inverse S T → Injection S T ``` * Added new proofs in `Function.Construct.Symmetry`: @@ -3276,7 +3325,6 @@ Additions to existing modules * Added new proof and record in `Function.Structures`: ```agda - IsLeftInverse.isSurjection : IsLeftInverse to from → IsSurjection to record IsSplitSurjection (f : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) ``` @@ -3285,6 +3333,14 @@ Additions to existing modules f⁻ = proj₁ ∘ surjective ``` +* Added new proof to `Induction.WellFounded` + ```agda + Acc-resp-flip-≈ : _<_ Respectsʳ (flip _≈_) → (Acc _<_) Respects _≈_ + acc⇒asym : Acc _<_ x → x < y → ¬ (y < x) + wf⇒asym : WellFounded _<_ → Asymmetric _<_ + wf⇒irrefl : _<_ Respects₂ _≈_ → Symmetric _≈_ → WellFounded _<_ → Irreflexive _≈_ _<_ + ``` + * Added new operations in `IO`: ``` Colist.forM : Colist A → (A → IO B) → IO (Colist B) @@ -3318,6 +3374,15 @@ Additions to existing modules prependVLams : List String → Term → Term ``` +* Added new types and operations to `Reflection.TCM`: + ``` + Blocker : Set + blockerMeta : Meta → Blocker + blockerAny : List Blocker → Blocker + blockerAll : List Blocker → Blocker + blockTC : Blocker → TC A + ``` + * Added new operations in `Relation.Binary.Construct.Closure.Equivalence`: ``` fold : IsEquivalence _∼_ → _⟶_ ⇒ _∼_ → EqClosure _⟶_ ⇒ _∼_ @@ -3343,21 +3408,21 @@ Additions to existing modules * Added new proofs to `Relation.Binary.Lattice.Properties.{Join,Meet}Semilattice`: ```agda - isPosemigroup : IsPosemigroup _≈_ _≤_ _∨_ - posemigroup : Posemigroup c ℓ₁ ℓ₂ - ≈-dec⇒≤-dec : Decidable _≈_ → Decidable _≤_ + 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 ℓ₁ ℓ₂ + commutativePomonoid : CommutativePomonoid c ℓ₁ ℓ₂ ``` * Added new proofs to `Relation.Binary.Properties.Poset`: ```agda - ≤-dec⇒≈-dec : Decidable _≤_ → Decidable _≈_ + ≤-dec⇒≈-dec : Decidable _≤_ → Decidable _≈_ ≤-dec⇒isDecPartialOrder : Decidable _≤_ → IsDecPartialOrder _≈_ _≤_ ``` @@ -3369,7 +3434,7 @@ Additions to existing modules * 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) + sym-cong : sym (cong f p) ≡ cong f (sym p) ``` * Added new proofs in `Relation.Binary.HeterogeneousEquality`: @@ -3388,49 +3453,51 @@ Additions to existing modules * 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 ℓ} _⊆′_ + ⊆-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 _⊆′_ _⊆′_ _⊆′_ + ⊆′-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 + U-irrelevant : Irrelevant U ∁-irrelevant : (P : Pred A ℓ) → Irrelevant (∁ P) ``` @@ -3453,9 +3520,9 @@ Additions to existing modules RightTrans R S = Trans R S R LeftTrans S R = Trans S R R - Dense _<_ = ∀ {x y} → x < y → ∃[ z ] x < z × z < y - Cotransitive _#_ = ∀ {x y} → x # y → ∀ z → (x # z) ⊎ (z # y) - Tight _≈_ _#_ = ∀ x y → (¬ x # y → x ≈ y) × (x ≈ y → ¬ x # y) + Dense _<_ = x < y → ∃[ z ] x < z × z < y + Cotransitive _#_ = x # y → ∀ z → (x # z) ⊎ (z # y) + Tight _≈_ _#_ = (¬ x # y → x ≈ y) × (x ≈ y → ¬ x # y) Monotonic₁ _≤_ _⊑_ f = f Preserves _≤_ ⟶ _⊑_ Antitonic₁ _≤_ _⊑_ f = f Preserves (flip _≤_) ⟶ _⊑_ @@ -3463,7 +3530,7 @@ Additions to existing modules 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) + Adjoint _≤_ _⊑_ f g = (f x ⊑ y → x ≤ g y) × (x ≤ g y → f x ⊑ y) ``` * Added new definitions in `Relation.Binary.Bundles`: @@ -3490,9 +3557,9 @@ Additions to existing modules * Added new proofs to `Relation.Binary.Construct.Closure.Transitive`: ``` - accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x + accessible⁻ : Acc _∼⁺_ x → Acc _∼_ x wellFounded⁻ : WellFounded _∼⁺_ → WellFounded _∼_ - accessible : ∀ {x} → Acc _∼_ x → Acc _∼⁺_ x + accessible : Acc _∼_ x → Acc _∼⁺_ x ``` * Added new operations in `Relation.Binary.PropositionalEquality.Properties`: @@ -3513,65 +3580,6 @@ Additions to existing modules 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 --------------------------------- @@ -3784,196 +3792,3 @@ This is a full list of proofs that have changed form to use irrelevant instance 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 proof, structure, and bundle to `Data.Rational.Properties` - ```agda - <-dense : Dense _<_ - <-isDenseLinearOrder : IsDenseLinearOrder _≡_ _<_ - <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ - ``` - -* 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) - - <-dense : Dense _<_ - ``` - -* Added new structures to `Data.Rational.Unnormalised.Properties` - ```agda - <-isDenseLinearOrder : IsDenseLinearOrder _≃_ _<_ - +-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ - +-*-isHeytingField : IsHeytingField _≃_ _≠_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ - ``` - -* Added new bundles to `Data.Rational.Unnormalised.Properties` - ```agda - <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ - +-*-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 _≈_ - - acc⇒asym : Acc _<_ x → x < y → ¬ (y < x) - wf⇒asym : WellFounded _<_ → Asymmetric _<_ - wf⇒irrefl : _<_ Respects₂ _≈_ → Symmetric _≈_ → - WellFounded _<_ → Irreflexive _≈_ _<_ - ``` - -* Added new types and operations to `Reflection.TCM`: - ``` - Blocker : Set - blockerMeta : Meta → Blocker - blockerAny : List Blocker → Blocker - blockerAll : List Blocker → Blocker - blockTC : Blocker → TC A - ``` - -* 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 ∎ - ``` From d2e94ef4e11e65d05e31f2db59b602859cea233b Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Fri, 20 Oct 2023 01:12:38 +0100 Subject: [PATCH 40/57] Spellcheck CHANGELOG (#2167) * spellcheck; `fix-whitespace`; unfixed: a few alignment issues; typo `predications` to `predicates` in `Relation.Unary.Relation.Binary.Equality` --- CHANGELOG.md | 430 +++++++++++++++++++++++++-------------------------- 1 file changed, 215 insertions(+), 215 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 29113663c0..535c899aaf 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -3,7 +3,7 @@ Version 2.0-rc1 The library has been tested using Agda 2.6.4. -NOTE: Version `2.0` contains various breaking changes and is not backwards +NOTE: Version `2.0` contains various breaking changes and is not backwards compatible with code written with version `1.X` of the library. Highlights @@ -34,8 +34,8 @@ Bug-fixes --------- * In `Algebra.Definitions.RawSemiring` the record `Prime` did not - enforce that the number was not divisible by `1#`. To fix this - `p∤1 : p ∤ 1#` hsa been added as a field. + enforce that the number was not divisible by `1#`. To fix this + `p∤1 : p ∤ 1#` has been added as a field. * 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 @@ -45,24 +45,24 @@ Bug-fixes * In `Data.Fin.Properties` the `i` argument to `opposite-suc` was implicit but could not be inferred in general. It has been made explicit. - -* In `Data.List.Membership.Setoid` the operations `_∷=_` and `_─_` + +* In `Data.List.Membership.Setoid` the operations `_∷=_` and `_─_` had an extraneous `{A : Set a}` parameter. This has been removed. -* In `Data.List.Relation.Ternary.Appending.Setoid` the constructors - were re-exported in their full generality which lead to unsolved meta +* In `Data.List.Relation.Ternary.Appending.Setoid` the constructors + were re-exported in their full generality which lead to unsolved meta variables at their use sites. Now versions of the constructors specialised to use the setoid's carrier set are re-exported. - -* In `Data.Nat.DivMod` the parameter `o` in the proof `/-monoˡ-≤` was + +* In `Data.Nat.DivMod` the parameter `o` in the proof `/-monoˡ-≤` was implicit but not inferrable. It has been changed to be explicit. -* In `Data.Nat.DivMod` the parameter `m` in the proof `+-distrib-/-∣ʳ` was +* In `Data.Nat.DivMod` the parameter `m` in the proof `+-distrib-/-∣ʳ` was implicit but not inferrable, while `n` is explicit but inferrable. They have been to explicit and implicit respectively. -* In `Data.Nat.GeneralisedArithmetic` the `s` and `z` arguments to the - following functions were implicit but no inferrable: +* In `Data.Nat.GeneralisedArithmetic` the `s` and `z` arguments to the + following functions were implicit but not inferrable: `fold-+`, `fold-k`, `fold-*`, `fold-pull`. They have been made explicit. * In `Data.Rational(.Unnormalised).Properties` the module `≤-Reasoning` @@ -79,9 +79,9 @@ Bug-fixes _≈⟨_⟨_ ↦ _≃⟨_⟨_ ``` -* In `Function.Construct.Composition` the combinators +* In `Function.Construct.Composition` the combinators `_⟶-∘_`, `_↣-∘_`, `_↠-∘_`, `_⤖-∘_`, `_⇔-∘_`, `_↩-∘_`, `_↪-∘_`, `_↔-∘_` - had their arguments in the wrong order. They have been flipped so they can + had their arguments in the wrong order. They have been flipped so they can actually be used as a composition operator. * In `Function.Definitions` the definitions of `Surjection`, `Inverseˡ`, @@ -98,7 +98,7 @@ Non-backwards compatible changes ### Removed deprecated names -* All modules and names that were deprecated in `v1.2` and before have +* All modules and names that were deprecated in `v1.2` and before have been removed. ### Changes to `LeftCancellative` and `RightCancellative` in `Algebra.Definitions` @@ -162,8 +162,8 @@ Non-backwards compatible changes `Algebra.Lattice.Properties.Semilattice` or `Algebra.Lattice.Morphism.Lattice`). See the `Deprecated modules` section below for full details. -* The definition of `IsDistributiveLattice` and `IsBooleanAlgebra` have changed so - that they are no longer right-biased which hindered compositionality. +* The definition of `IsDistributiveLattice` and `IsBooleanAlgebra` have changed so + that they are no longer right-biased which hindered compositionality. More concretely, `IsDistributiveLattice` now has fields: ```agda ∨-distrib-∧ : ∨ DistributesOver ∧ @@ -225,12 +225,12 @@ Non-backwards compatible changes IsSemiringHomomorphism IsRingHomomorphism ``` - all had two separate proofs of `IsRelHomomorphism` within them. - -* To fix this they have all been redefined to build up from `IsMonoidHomomorphism`, - `IsNearSemiringHomomorphism`, and `IsSemiringHomomorphism` respectively, - adding a single property at each step. - + all had two separate proofs of `IsRelHomomorphism` within them. + +* To fix this they have all been redefined to build up from `IsMonoidHomomorphism`, + `IsNearSemiringHomomorphism`, and `IsSemiringHomomorphism` respectively, + adding a single property at each step. + * Similarly, `IsLatticeHomomorphism` is now built as `IsRelHomomorphism` along with proofs that `_∧_` and `_∨_` are homomorphic. @@ -241,13 +241,13 @@ Non-backwards compatible changes * 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 + 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. * Full list of moved modules: @@ -284,7 +284,7 @@ Non-backwards compatible changes IO.Categorical ↦ IO.Effectful Reflection.TCM.Categorical ↦ Reflection.TCM.Effectful ``` - + * Full list of new modules: ``` Algebra.Construct.Initial @@ -374,10 +374,10 @@ Non-backwards compatible changes * 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 `Codata.Sized`, e.g. `Codata.Stream` + will be needed using coinductive records. + +* Therefore all existing modules in `Codata` which used sized types have been + moved inside a new folder named `Codata.Sized`, e.g. `Codata.Stream` has become `Codata.Sized.Stream`. ### New proof-irrelevant for empty type in `Data.Empty` @@ -398,14 +398,14 @@ Non-backwards compatible changes + 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 ``` @@ -424,7 +424,7 @@ Non-backwards compatible changes ≺-rec ``` these functions are also deprecated. - + * Likewise in `Data.Fin.Properties` the proofs `≺⇒<′` and `<′⇒≺` have been deprecated in favour of their proxy counterparts `<⇒<′` and `<′⇒<`. @@ -473,9 +473,9 @@ Non-backwards compatible changes #### Standardisation of `lookup` in `Data.(List/Vec/...)` -* All the types of `lookup` functions (and variants) in the following modules - have been changed to match the argument convention adopted in the `List` module (i.e. - `lookup` takes its container first and the index, whose type may depend on the +* All the types of `lookup` functions (and variants) in the following modules + have been changed to match the argument convention adopted in the `List` module (i.e. + `lookup` takes its container first and the index, whose type may depend on the container value, second): ``` Codata.Guarded.Stream @@ -500,7 +500,7 @@ Non-backwards compatible changes Data.Vec.Recursive ``` -* To accomodate this in in `Data.Vec.Relation.Unary.Linked.Properties` +* To accommodate this in in `Data.Vec.Relation.Unary.Linked.Properties` and `Codata.Guarded.Stream.Relation.Binary.Pointwise`, the proofs called `lookup` have been renamed `lookup⁺`. @@ -509,11 +509,11 @@ Non-backwards compatible changes * 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 + 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 + +* 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 discussion for a fuller explanation of the motivation and implementation. @@ -538,47 +538,47 @@ Non-backwards compatible changes * 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. - ```agda - 0/n≡0 : 0 / suc n ≡ 0 - ``` - - 2. *Add an instance argument parameter* - 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. + 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. + ```agda + 0/n≡0 : 0 / suc n ≡ 0 + ``` + + 2. *Add an instance argument parameter* - 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. ```agda 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. - ```agda - instance - n≢0 : NonZero n - 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. + ```agda + instance + n≢0 : NonZero n + n≢0 = ... - 0/n≡0 : 0 / 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 - ``` + 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 +* 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 + 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 @@ -588,13 +588,13 @@ Non-backwards compatible changes ``` 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 + 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, in `Data.Rational.Unnormalised.Base` the definition of `_≢0` is now - redundent and so has been removed. Additionally the following proofs about it have + redundant and so has been removed. Additionally the following proofs about it have also been removed from `Data.Rational.Unnormalised.Properties`: ``` p≄0⇒∣↥p∣≢0 : ∀ p → p ≠ 0ℚᵘ → ℤ.∣ (↥ p) ∣ ≢0 @@ -613,9 +613,9 @@ Non-backwards compatible changes ``` which introduced a spurious additional definition, when this is in fact, modulo field names and implicit/explicit qualifiers, equivalent to the definition of left- - divisibility, `_∣ˡ_` for the `RawMagma` structure of `_+_`. - -* Since the addition of raw bundles to `Data.X.Base`, this definition can now be + divisibility, `_∣ˡ_` for the `RawMagma` structure of `_+_`. + +* Since the addition of raw bundles to `Data.X.Base`, this definition can now be made directly. Accordingly, the definition has been changed to: ```agda _≤″_ : (m n : ℕ) → Set @@ -625,8 +625,8 @@ Non-backwards compatible changes ``` * Knock-on consequences include the need to retain the old constructor - name, now introduced as a pattern synonym, and introduction of (a function - equivalent to) the former field name/projection function `proof` as + name, now introduced as a pattern synonym, and introduction of (a function + equivalent to) the former field name/projection function `proof` as `≤″-proof` in `Data.Nat.Properties`. ### Changes to definition of `Prime` in `Data.Nat.Primality` @@ -672,24 +672,24 @@ Non-backwards compatible changes ``` * As a consequence of this, some proofs that relied either on this reduction behaviour - or on eta-equality may no longer type-check. - + or on eta-equality may no longer type-check. + * 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. * Similarly, in order to prevent reduction, the equality `_≃_` in `Data.Rational.Base` - has been made into a data type with the single constructor `*≡*`. The destructor + has been made into a data type with the single constructor `*≡*`. The destructor `drop-*≡*` has been added to `Data.Rational.Properties`. #### Deprecation of old `Function` hierarchy @@ -697,7 +697,7 @@ Non-backwards compatible changes * The new `Function` hierarchy was introduced in `v1.2` which follows the same `Core`/`Definitions`/`Bundles`/`Structures` as all the other hierarchies in the library. - + * At the time the old function hierarchy in: ``` Function.Equality @@ -711,11 +711,11 @@ Non-backwards compatible changes ``` was unofficially deprecated, but could not be officially deprecated because of it's widespread use in the rest of the library. - -* Now, the old hierarchy modules have all been officially deprecated. All + +* Now, the old hierarchy modules have all been officially deprecated. All uses of them in the rest of the library have been switched to use the - new hierarachy. - + new hierarchy. + * The latter is unfortunately a relatively big breaking change, but was judged to be unavoidable given how widely used the old hierarchy was. @@ -724,9 +724,9 @@ Non-backwards compatible changes * In `Function.Bundles` the names of the fields in the records have been changed from `f`, `g`, `cong₁` and `cong₂` to `to`, `from`, `to-cong`, `from-cong`. -* In `Function.Definitions`, the module 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 two modules have +* In `Function.Definitions`, the module 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 two modules have been removed and their definitions folded into `Function.Definitions`. * In `Function.Definitions` the following definitions have been changed from: @@ -741,9 +741,9 @@ Non-backwards compatible changes 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. + 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`, @@ -769,9 +769,9 @@ Non-backwards compatible changes WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _ WfRec _<_ P x = ∀ y → y < x → P y ``` - which meant that all arguments involving accessibility and wellfoundedness proofs - were polluted by almost-always-inferrable explicit arguments for the `y` position. - + which meant that all arguments involving accessibility and wellfoundedness proofs + were polluted by almost-always-inferrable explicit arguments for the `y` position. + * The definition has now been changed to make that argument *implicit*, as ```agda WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _ @@ -821,44 +821,44 @@ Non-backwards compatible changes ### 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 about these definitions were spread over +* 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 about these definitions were spread over `Relation.Nullary.(Negation/Product/Sum/Implication etc.)`. * To fix this all the contents of the latter is now exported by `Relation.Nullary`. * In order to achieve this the following backwards compatible changes have been made: - + 1. the definition of `Dec` and `recompute` have been moved to `Relation.Nullary.Decidable.Core` - + 2. the definition of `Reflects` has been moved to `Relation.Nullary.Reflects` - + 3. the definition of `¬_` has been moved to `Relation.Nullary.Negation.Core` - 4. The modules `Relation.Nullary.(Product/Sum/Implication)` have been deprecated - and their contents moved to `Relation.Nullary.(Negation/Reflects/Decidable)`. + 4. The modules `Relation.Nullary.(Product/Sum/Implication)` have been deprecated + and their contents moved to `Relation.Nullary.(Negation/Reflects/Decidable)`. as well as the following breaking changes: - 1. `¬?` has been moved from `Relation.Nullary.Negation.Core` to - `Relation.Nullary.Decidable.Core` - - 2. `¬-reflects` has been moved from `Relation.Nullary.Negation.Core` to - `Relation.Nullary.Reflects`. - - 3. `decidable-stable`, `excluded-middle` and `¬-drop-Dec` have been moved - from `Relation.Nullary.Negation` to `Relation.Nullary.Decidable`. - + 1. `¬?` has been moved from `Relation.Nullary.Negation.Core` to + `Relation.Nullary.Decidable.Core` + + 2. `¬-reflects` has been moved from `Relation.Nullary.Negation.Core` to + `Relation.Nullary.Reflects`. + + 3. `decidable-stable`, `excluded-middle` and `¬-drop-Dec` have been moved + from `Relation.Nullary.Negation` to `Relation.Nullary.Decidable`. + 4. `fromDec` and `toDec` have been moved from `Data.Sum.Base` to `Data.Sum`. ### (Issue #2096) Introduction of flipped and negated relation symbols to bundles in `Relation.Binary.Bundles` -* Previously, bundles such as `Preorder`, `Poset`, `TotalOrder` etc. did not have the flipped +* Previously, bundles such as `Preorder`, `Poset`, `TotalOrder` etc. did not have the flipped and negated versions of the operators exposed. In some cases they could obtained by opening the relevant `Relation.Binary.Properties.X` file but usually they had to be redefined every time. - + * To fix this, these bundles now all export all 4 versions of the operator: normal, converse, negated, converse-negated. Accordingly they are no longer exported from the corresponding `Properties` file. @@ -867,26 +867,26 @@ Non-backwards compatible changes converse relation could only be discussed *semantically* in terms of `flip _∼_`. * Now, the `Preorder` record field `_∼_` has been renamed to `_≲_`, with `_≳_` - introduced as a definition in `Relation.Binary.Bundles.Preorder`. - Partial backwards compatible has been achieved by redeclaring a deprecated version - of the old symbol in the record. Therefore, only _declarations_ of `PartialOrder` records will + introduced as a definition in `Relation.Binary.Bundles.Preorder`. + Partial backwards compatible has been achieved by redeclaring a deprecated version + of the old symbol in the record. Therefore, only _declarations_ of `PartialOrder` records will need their field names updating. ### Changes to definition of `IsStrictTotalOrder` in `Relation.Binary.Structures` -* The previous definition of the record `IsStrictTotalOrder` did not - build upon `IsStrictPartialOrder` as would be expected. +* The previous definition of the record `IsStrictTotalOrder` did not + build upon `IsStrictPartialOrder` as would be expected. Instead it omitted several fields like irreflexivity as they were derivable from the - proof of trichotomy. However, this led to problems further up the hierarchy where - bundles such as `StrictTotalOrder` which contained multiple distinct proofs of + proof of trichotomy. However, this led to problems further up the hierarchy where + bundles such as `StrictTotalOrder` which contained multiple distinct proofs of `IsStrictPartialOrder`. -* To remedy this the definition of `IsStrictTotalOrder` has been changed to so +* To remedy this the definition of `IsStrictTotalOrder` has been changed to so that it builds upon `IsStrictPartialOrder` as would be expected. -* To aid migration, the old record definition has been moved to - `Relation.Binary.Structures.Biased` which contains the `isStrictTotalOrderᶜ` +* To aid migration, the old record definition has been moved to + `Relation.Binary.Structures.Biased` which contains the `isStrictTotalOrderᶜ` smart constructor (which is re-exported by `Relation.Binary`) . Therefore the old code: ```agda <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ @@ -917,14 +917,14 @@ Non-backwards compatible changes ### Changes to the interface of `Relation.Binary.Reasoning.Triple` -* The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof +* The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof that the strict relation is irreflexive. * This allows the addition of 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 + that takes a proof that a value is strictly less than itself and then applies the principle of explosion to derive anything. * Specialised versions of this combinator are available in the `Reasoning` modules @@ -945,9 +945,9 @@ Non-backwards compatible changes * Previously, every `Reasoning` module in the library tended to roll it's own set of syntax for the combinators. This hindered consistency and maintainability. - -* To improve the situation, a new module `Relation.Binary.Reasoning.Syntax` - has been introduced which exports a wide range of sub-modules containing + +* To improve the situation, a new module `Relation.Binary.Reasoning.Syntax` + has been introduced which exports a wide range of sub-modules containing pre-existing reasoning combinator syntax. * This makes it possible to add new or rename existing reasoning combinators to a @@ -956,39 +956,39 @@ Non-backwards compatible changes * One pre-requisite for that is that `≡-Reasoning` has been moved from `Relation.Binary.PropositionalEquality.Core` (which shouldn't be - imported anyway as it's a `Core` module) to + imported anyway as it's a `Core` module) to `Relation.Binary.PropositionalEquality.Properties`. It is still exported by `Relation.Binary.PropositionalEquality`. ### Renaming of symmetric combinators in `Reasoning` modules -* We've had various complaints about the symmetric version of reasoning combinators +* We've had various complaints about the symmetric version of reasoning combinators that use the syntax `_R˘⟨_⟩_` for some relation `R`, (e.g. `_≡˘⟨_⟩_` and `_≃˘⟨_⟩_`) introduced in `v1.0`. In particular: 1. The symbol `˘` is hard to type. 2. The symbol `˘` doesn't convey the direction of the equality 3. The left brackets aren't vertically aligned with the left brackets of the non-symmetric version. - -* To address these problems we have renamed all the symmetric versions of the + +* To address these problems we have renamed all the symmetric versions of the combinators from `_R˘⟨_⟩_` to `_R⟨_⟨_` (the direction of the last bracket is flipped to indicate the quality goes from right to left). - -* The old combinators still exist but have been deprecated. However due to + +* The old combinators still exist but have been deprecated. However due to [Agda issue #5617](https://github.com/agda/agda/issues/5617), the deprecation warnings don't fire correctly. We will not remove the old combinators before the above issue is addressed. However, we still encourage migrating to the new combinators! * On a Linux-based system, the following command was used to globally migrate all uses of the - old combinators to the new ones in the standard library itself. + old combinators to the new ones in the standard library itself. It *may* be useful when trying to migrate your own code: ```bash find . -type f -name "*.agda" -print0 | xargs -0 sed -i -e 's/˘⟨\(.*\)⟩/⟨\1⟨/g' ``` - USUAL CAVEATS: It has not been widely tested and the standard library developers are not - responsible for the results of this command. It is strongly recommended you back up your + USUAL CAVEATS: It has not been widely tested and the standard library developers are not + responsible for the results of this command. It is strongly recommended you back up your work before you attempt to run it. - -* NOTE: this refactoring may require some extra bracketing around the operator `_⟨_⟩_` from + +* NOTE: this refactoring may require some extra bracketing around the operator `_⟨_⟩_` from `Function.Base` if the `_⟨_⟩_` operator is used within the reasoning proof. The symptom for this is a `Don't know how to parse` error. @@ -1035,10 +1035,10 @@ Non-backwards compatible changes * `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. + 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. * In `Data.Fin.Base` the relations `_≤_` `_≥_` `_<_` `_>_` have been @@ -1047,9 +1047,9 @@ Non-backwards compatible changes properties about the orderings themselves the second index must be provided explicitly. -* In `Data.Fin.Properties` the proof `inj⇒≟` 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 +* In `Data.Fin.Properties` the proof `inj⇒≟` that an injection from a type + `A` into `Fin n` induces a decision procedure for `_≡_` on `A` has been + generalised to other equivalences over `A` (i.e. to arbitrary setoids), and renamed from `eq?` to the more descriptive and `inj⇒decSetoid`. * In `Data.Fin.Properties` the proof `pigeonhole` has been strengthened @@ -1058,19 +1058,19 @@ Non-backwards compatible changes * In `Data.Fin.Substitution.TermSubst` the fixity of `_/Var_` has been changed from `infix 8` to `infixl 8`. -* In `Data.Integer.DivMod` the previous implementations of - `_divℕ_`, `_div_`, `_modℕ_`, `_mod_` internally converted to the unary - `Fin` datatype resulting in poor performance. The implementation has been +* In `Data.Integer.DivMod` the previous implementations of + `_divℕ_`, `_div_`, `_modℕ_`, `_mod_` 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. -* In `Data.Integer.Properties` the first two arguments of `m≡n⇒m-n≡0` +* In `Data.Integer.Properties` the first two arguments of `m≡n⇒m-n≡0` (now renamed `i≡j⇒i-j≡0`) have been made implicit. -* In `Data.(List|Vec).Relation.Binary.Lex.Strict` the argument `xs` in +* In `Data.(List|Vec).Relation.Binary.Lex.Strict` the argument `xs` in `xs≮[]` in introduced in PRs #1648 and #1672 has now been made implicit. -* In `Data.List.NonEmpty` the functions `split`, `flatten` and `flatten-split` +* In `Data.List.NonEmpty` the functions `split`, `flatten` and `flatten-split` have been removed from. 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 @@ -1082,20 +1082,20 @@ Non-backwards compatible changes proofs. * In `Data.Nat.Divisibility` the proof `m/n/o≡m/[n*o]` has been removed. In it's - place a new more general proof `m/n/o≡m/[n*o]` has been added to `Data.Nat.DivMod` + place a new more general proof `m/n/o≡m/[n*o]` has been added to `Data.Nat.DivMod` that doesn't require the `n * o ∣ m` pre-condition. -* In `Data.Product.Relation.Binary.Lex.Strict` the proof of wellfoundedness +* In `Data.Product.Relation.Binary.Lex.Strict` the proof of wellfoundedness of the lexicographic ordering on products, no longer requires the assumption of symmetry for the first equality relation `_≈₁_`. -* In `Data.Rational.Base` the constructors `+0` and `+[1+_]` from `Data.Integer.Base` +* In `Data.Rational.Base` the constructors `+0` and `+[1+_]` from `Data.Integer.Base` are no longer re-exported by `Data.Rational.Base`. You will have to open `Data.Integer(.Base)` directly to use them. -* In `Data.Rational(.Unnormalised).Properties` the types of the proofs - `pos⇒1/pos`/`1/pos⇒pos` and `neg⇒1/neg`/`1/neg⇒neg` have been switched, - as the previous naming scheme didn't correctly generalise to e.g. `pos+pos⇒pos`. +* In `Data.Rational(.Unnormalised).Properties` the types of the proofs + `pos⇒1/pos`/`1/pos⇒pos` and `neg⇒1/neg`/`1/neg⇒neg` 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: ```agda pos⇒1/pos : ∀ p .{{_ : NonZero p}} .{{Positive (1/ p)}} → Positive p @@ -1111,7 +1111,7 @@ Non-backwards compatible changes * In `Data.Vec.Base` the definition of `_>>=_` under `Data.Vec.Base` has been moved to the submodule `CartesianBind` in order to avoid clashing with the - new, alternative definition of `_>>=_`, located in the second new submodule + new, alternative definition of `_>>=_`, located in the second new submodule `DiagonalBind`. * In `Data.Vec.Base` the definitions `init` and `last` have been changed from the `initLast` @@ -1119,15 +1119,15 @@ Non-backwards compatible changes * In `Data.Vec.Properties` the type of the proof `zipWith-comm` has been generalised from: ```agda - zipWith-comm : ∀ {f : A → A → B} (comm : ∀ x y → f x y ≡ f y x) (xs ys : Vec A n) → + zipWith-comm : ∀ {f : A → A → B} (comm : ∀ x y → f x y ≡ f y x) (xs ys : Vec A n) → zipWith f xs ys ≡ zipWith f ys xs ``` to ```agda - zipWith-comm : ∀ {f : A → B → C} {g : B → A → C} (comm : ∀ x y → f x y ≡ g y x) (xs : Vec A n) ys → + zipWith-comm : ∀ {f : A → B → C} {g : B → A → C} (comm : ∀ x y → f x y ≡ g y x) (xs : Vec A n) ys → zipWith f xs ys ≡ zipWith g ys xs ``` - + * In `Data.Vec.Relation.Unary.All` the functions `lookup` and `tabulate` have been moved to `Data.Vec.Relation.Unary.All.Properties` and renamed `lookup⁺` and `lookup⁻` respectively. @@ -1139,15 +1139,15 @@ Non-backwards compatible changes iterate : (A → A) → A → ∀ n → Vec A n replicate : ∀ n → A → Vec A n ``` - -* In `Relation.Binary.Construct.Closure.Symmetric` the operation + +* In `Relation.Binary.Construct.Closure.Symmetric` the operation `SymClosure` on relations in 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 `Relation.Nullary.Decidable.Core` the name `excluded-middle` has been +* In `Relation.Nullary.Decidable.Core` the name `excluded-middle` has been renamed to `¬¬-excluded-middle`. @@ -1185,10 +1185,10 @@ Other major improvements * Raw bundles by design don't contain any proofs so should in theory be able to live in `Data.X.Base` rather than `Data.X.Bundles`. - + * To achieve this while keeping the dependency footprint small, the definitions of - raw bundles (`RawMagma`, `RawMonoid` etc.) have been moved from `Algebra(.Lattice)Bundles` to - a new module `Algebra(.Lattice).Bundles.Raw` which can be imported at much lower cost + raw bundles (`RawMagma`, `RawMonoid` etc.) have been moved from `Algebra(.Lattice)Bundles` to + a new module `Algebra(.Lattice).Bundles.Raw` which can be imported at much lower cost from `Data.X.Base`. * We have then moved raw bundles defined in `Data.X.Properties` to `Data.X.Base` for @@ -1201,7 +1201,7 @@ Deprecated modules * This fixes the fact we had picked the wrong name originally. The erased modality corresponds to `@0` whereas the irrelevance one corresponds to `.`. - + ### Deprecation of `Data.Fin.Substitution.Example` * The module `Data.Fin.Substitution.Example` has been deprecated, and moved to `README.Data.Fin.Substitution.UntypedLambda` @@ -1212,7 +1212,7 @@ Deprecated modules ### Deprecation of `Data.Product.Function.Dependent.Setoid.WithK` -* This module has been deprecated, as none of its contents actually depended on axiom K. +* 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`. ### Moving `Function.Related` @@ -1444,7 +1444,7 @@ Deprecated names * In `Data.List.Relation.Unary.All.Properties`: ```agda gmap ↦ gmap⁺ - + updateAt-id-relative ↦ updateAt-id-local updateAt-compose-relative ↦ updateAt-updateAt-local updateAt-compose ↦ updateAt-updateAt @@ -2045,7 +2045,7 @@ New modules Relation.Unary.Algebra ``` -* Both versions of equality on predications are equivalences +* Both versions of equality on predicates are equivalences ``` Relation.Unary.Relation.Binary.Equality ``` @@ -2195,8 +2195,8 @@ Additions to existing modules SemiringWithoutAnnihilatingZero b ℓ₂ → SemiringWithoutAnnihilatingZero (a ⊔ b) (ℓ₁ ⊔ ℓ₂) semiring : Semiring a ℓ₁ → Semiring b ℓ₂ → Semiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) - commutativeSemiring : CommutativeSemiring a ℓ₁ → - CommutativeSemiring 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) (ℓ₁ ⊔ ℓ₂) @@ -2204,14 +2204,14 @@ Additions to existing modules 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 : 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) (ℓ₁ ⊔ ℓ₂) + 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) (ℓ₁ ⊔ ℓ₂) @@ -2224,9 +2224,9 @@ Additions to existing modules 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) (ℓ₁ ⊔ ℓ₂) + nonAssociativeRing : NonAssociativeRing a ℓ₁ → + NonAssociativeRing b ℓ₂ → + NonAssociativeRing (a ⊔ b) (ℓ₁ ⊔ ℓ₂) quasiring : Quasiring a ℓ₁ → Quasiring b ℓ₂ → Quasiring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) nearring : Nearring a ℓ₁ → Nearring b ℓ₂ → Nearring (a ⊔ b) (ℓ₁ ⊔ ℓ₂) ``` @@ -2308,7 +2308,7 @@ Additions to existing modules * Added new proofs to `Algebra.Properties.CommutativeSemigroup`: ```agda xy∙xx≈x∙yxx : (x ∙ y) ∙ (x ∙ x) ≈ x ∙ (y ∙ (x ∙ x)) - + interchange : Interchangable _∙_ _∙_ leftSemimedial : LeftSemimedial _∙_ rightSemimedial : RightSemimedial _∙_ @@ -2406,7 +2406,7 @@ Additions to existing modules xor-annihilates-not : (not x) xor (not y) ≡ x xor y ``` -* Exposed container combinator conversion functions from `Data.Container.Combinator.Properties` +* Exposed container combinator conversion functions from `Data.Container.Combinator.Properties` separately from their correctness proofs in `Data.Container.Combinator`: ``` to-id : F.id A → ⟦ id ⟧ A @@ -2437,7 +2437,7 @@ Additions to existing modules ```agda spo-wellFounded : IsStrictPartialOrder _≈_ _⊏_ → WellFounded _⊏_ spo-noetherian : IsStrictPartialOrder _≈_ _⊏_ → WellFounded (flip _⊏_) - + <-weakInduction-startingFrom : P i → (∀ j → P (inject₁ j) → P (suc j)) → ∀ {j} → j ≥ i → P j ``` @@ -2508,7 +2508,7 @@ Additions to existing modules sub≡sub : map var (VarSubst.sub x) ≡ T.sub (T.var x) ↑≡↑ : map var (ρ VarSubst.↑) ≡ map T.var ρ T.↑ /Var≡/ : t /Var ρ ≡ t T./ map T.var ρ - + sub-renaming-commutes : t /Var VarSubst.sub x T./ ρ ≡ t T./ ρ T.↑ T./ T.sub (lookup ρ x) sub-commutes-with-renaming : t T./ T.sub t′ /Var ρ ≡ t /Var ρ VarSubst.↑ T./ T.sub (t′ /Var ρ) ``` @@ -2602,9 +2602,9 @@ Additions to existing modules * Added new proofs to `Data.List.Relation.Binary.Sublist.Setoid.Properties`: ``` ⊆-mergeˡ : xs ⊆ merge _≤?_ xs ys - ⊆-mergeʳ : ys ⊆ merge _≤?_ xs ys + ⊆-mergeʳ : ys ⊆ merge _≤?_ xs ys ``` - + * Added new functions in `Data.List.Relation.Unary.All`: ``` decide : Π[ P ∪ Q ] → Π[ All P ∪ Any Q ] @@ -2631,7 +2631,7 @@ Additions to existing modules 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₂ - + ∈-++⁺∘++⁻ : (p : v ∈ xs ++ ys) → [ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ (∈-++⁻ xs p) ≡ p ∈-++⁻∘++⁺ : (p : v ∈ xs ⊎ v ∈ ys) → ∈-++⁻ xs ([ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ p) ≡ p ∈-++↔ : (v ∈ xs ⊎ v ∈ ys) ↔ v ∈ xs ++ ys @@ -2682,12 +2682,12 @@ Additions to existing modules length-removeAt′ : length xs ≡ suc (length (removeAt xs k)) removeAt-insertAt : removeAt (insertAt xs i v) ((cast (sym (length-insertAt xs i v)) i)) ≡ xs insertAt-removeAt : insertAt (removeAt xs i) (cast (sym (lengthAt-removeAt xs i)) i) (lookup xs i) ≡ xs - + 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 ``` @@ -2830,7 +2830,7 @@ Additions to existing modules ⊔≡⊔′ : m ⊔ n ≡ m ⊔′ n ⊓≡⊓′ : m ⊓ n ≡ m ⊓′ n ∣-∣≡∣-∣′ : ∣ m - n ∣ ≡ ∣ m - n ∣′ - + nonZero? : Decidable NonZero eq? : A ↣ ℕ → DecidableEquality A ≤-<-connex : Connex _≤_ _<_ @@ -2943,7 +2943,7 @@ Additions to existing modules pos⇒nonZero : .{{Positive p}} → NonZero p neg⇒nonZero : .{{Negative p}} → NonZero p nonZero⇒1/nonZero : .{{_ : NonZero p}} → NonZero (1/ p) - + <-dense : Dense _<_ <-isDenseLinearOrder : IsDenseLinearOrder _≡_ _<_ <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ @@ -2984,7 +2984,7 @@ Additions to existing modules 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) - + 0≄1 : 0ℚᵘ ≄ 1ℚᵘ ≃-≄-irreflexive : Irreflexive _≃_ _≄_ ≄-symmetric : Symmetric _≄_ @@ -2992,15 +2992,15 @@ Additions to existing modules ≄⇒invertible : p ≄ q → Invertible _≃_ 1ℚᵘ _*_ (p - q) <-dense : Dense _<_ - + <-isDenseLinearOrder : IsDenseLinearOrder _≃_ _<_ +-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ +-*-isHeytingField : IsHeytingField _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ - + <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ +-*-heytingCommutativeRing : HeytingCommutativeRing 0ℓ 0ℓ 0ℓ +-*-heytingField : HeytingField 0ℓ 0ℓ 0ℓ - + module ≃-Reasoning = SetoidReasoning ≃-setoid ``` @@ -3228,7 +3228,7 @@ Additions to existing modules * Added new proofs in `Data.Vec.Relation.Binary.Lex.Strict`: ```agda xs≮[] : ¬ xs < [] - + <-respectsˡ : IsPartialEquivalence _≈_ → _≺_ Respectsˡ _≈_ → _<_ Respectsˡ _≋_ <-respectsʳ : IsPartialEquivalence _≈_ → _≺_ Respectsʳ _≈_ → _<_ _Respectsʳ _≋_ <-wellFounded : Transitive _≈_ → _≺_ Respectsʳ _≈_ → WellFounded _≺_ → WellFounded _<_ @@ -3252,7 +3252,7 @@ Additions to existing modules * Added new functions in `Data.Vec.Relation.Unary.All`: ``` decide : Π[ P ∪ Q ] → Π[ All P ∪ Any Q ] - + 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) @@ -3263,7 +3263,7 @@ Additions to existing modules ``` ++-assoc : (xs ++ ys) ++ zs ≋ xs ++ (ys ++ zs) ``` - + * Added new isomorphisms to `Data.Vec.N-ary`: ```agda Vec↔N-ary : ∀ n → (Vec A n → B) ↔ N-ary n A B @@ -3281,7 +3281,7 @@ Additions to existing modules evalState : State s a → s → a execState : State s a → s → s ``` - + * Added a non-dependent version of `flip` in `Function.Base`: ```agda flip′ : (A → B → C) → (B → A → C) @@ -3300,10 +3300,10 @@ Additions to existing modules ↔-refl : Reflexive _↔_ ↔-sym : Symmetric _↔_ ↔-trans : Transitive _↔_ - + ↔⇒↣ : A ↔ B → A ↣ B ↔-fun : A ↔ B → C ↔ D → (A → C) ↔ (B → D) - + Inverse⇒Injection : Inverse S T → Injection S T ``` @@ -3322,7 +3322,7 @@ Additions to existing modules _!|>_ : (a : A) → (∀ a → B a) → B a _!|>′_ : A → (A → B) → B ``` - + * Added new proof and record in `Function.Structures`: ```agda record IsSplitSurjection (f : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) @@ -3465,7 +3465,7 @@ Additions to existing modules ⊂-resp-≐ : _⊂_ Respects₂ _≐_ ⊂-irrefl : Irreflexive _≐_ _⊂_ ⊂-antisym : Antisymmetric _≐_ _⊂_ - + ∅-⊆′ : (P : Pred A ℓ) → ∅ ⊆′ P ⊆′-U : (P : Pred A ℓ) → P ⊆′ U ⊆′-refl : Reflexive {A = Pred A ℓ} _⊆′_ @@ -3487,7 +3487,7 @@ Additions to existing modules ⊆′⇒⊆ : _⊆′_ ⇒ _⊆_ ⊂⇒⊂′ : _⊂_ ⇒ _⊂′_ ⊂′⇒⊂ : _⊂′_ ⇒ _⊂_ - + ≐-refl : Reflexive _≐_ ≐-sym : Sym _≐_ _≐_ ≐-trans : Trans _≐_ _≐_ _≐_ From 0befc0dedf6abf50d7f4a96af10dac7280816799 Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Sat, 21 Oct 2023 08:37:35 +0100 Subject: [PATCH 41/57] Fixed Agda version typo in README (#2176) --- README.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/README.agda b/README.agda index ced0f11596..cc436d0848 100644 --- a/README.agda +++ b/README.agda @@ -19,7 +19,7 @@ module README where -- and other anonymous contributors. ------------------------------------------------------------------------ --- This version of the library has been tested using Agda 2.6.3. +-- This version of the library has been tested using Agda 2.6.4. -- The library comes with a .agda-lib file, for use with the library -- management system. From cd34417723e2f5eff24eb19a1a409732e3cb549e Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Sat, 21 Oct 2023 08:38:19 +0100 Subject: [PATCH 42/57] =?UTF-8?q?Fixed=20in=20deprecation=20warning=20for?= =?UTF-8?q?=20`<-trans=CB=A1`=20(#2173)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/Data/Nat/Properties.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda index e304ed659f..37b1f33ee0 100644 --- a/src/Data/Nat/Properties.agda +++ b/src/Data/Nat/Properties.agda @@ -2342,6 +2342,6 @@ open Data.Nat.Base public <-transˡ = <-≤-trans {-# WARNING_ON_USAGE <-transˡ -"Warning: <-transˡ was deprecated in v2.0. Please use ≤-<-trans instead. " +"Warning: <-transˡ was deprecated in v2.0. Please use <-≤-trans instead. " #-} From 496ededbed9c037d8a57927e5e7a95794df74f26 Mon Sep 17 00:00:00 2001 From: Andreas Abel Date: Wed, 1 Nov 2023 03:31:40 +0100 Subject: [PATCH 43/57] Bump Haskell CI (original!) to latest minor GHC versions (#2187) --- .github/workflows/ci-ubuntu.yml | 14 ++++++------ .github/workflows/haskell-ci.yml | 38 +++++++++----------------------- agda-stdlib-utils.cabal | 6 ++--- 3 files changed, 21 insertions(+), 37 deletions(-) diff --git a/.github/workflows/ci-ubuntu.yml b/.github/workflows/ci-ubuntu.yml index 2ecc969589..51c5f7b09c 100644 --- a/.github/workflows/ci-ubuntu.yml +++ b/.github/workflows/ci-ubuntu.yml @@ -8,6 +8,7 @@ on: branches: - master - experimental + merge_group: ######################################################################## ## CONFIGURATION @@ -71,16 +72,15 @@ jobs: - name: Initialise variables run: | - if [[ '${{ github.ref }}' == 'refs/heads/master' \ - || '${{ github.base_ref }}' == 'master' ]]; then - # Pick Agda version for master - echo "AGDA_COMMIT=tags/v2.6.4" >> $GITHUB_ENV; - echo "AGDA_HTML_DIR=html/master" >> $GITHUB_ENV - elif [[ '${{ github.ref }}' == 'refs/heads/experimental' \ - || '${{ github.base_ref }}' == 'experimental' ]]; then + if [[ '${{ github.ref }}' == 'refs/heads/experimental' \ + || '${{ github.base_ref }}' == 'experimental' ]]; then # Pick Agda version for experimental echo "AGDA_COMMIT=4d36cb37f8bfb765339b808b13356d760aa6f0ec" >> $GITHUB_ENV; echo "AGDA_HTML_DIR=html/experimental" >> $GITHUB_ENV + else + # Pick Agda version for master + echo "AGDA_COMMIT=tags/v2.6.4" >> $GITHUB_ENV; + echo "AGDA_HTML_DIR=html/master" >> $GITHUB_ENV fi if [[ '${{ github.ref }}' == 'refs/heads/master' \ diff --git a/.github/workflows/haskell-ci.yml b/.github/workflows/haskell-ci.yml index b8c55c19be..5f9fa86b50 100644 --- a/.github/workflows/haskell-ci.yml +++ b/.github/workflows/haskell-ci.yml @@ -8,9 +8,9 @@ # # For more information, see https://github.com/haskell-CI/haskell-ci # -# version: 0.17.20230824 +# version: 0.17.20231010 # -# REGENDATA ("0.17.20230824",["github","--no-cabal-check","agda-stdlib-utils.cabal"]) +# REGENDATA ("0.17.20231010",["github","--no-cabal-check","agda-stdlib-utils.cabal"]) # name: Haskell-CI on: @@ -32,6 +32,7 @@ on: - agda-stdlib-utils.cabal - cabal.haskell-ci - "*.hs" + merge_group: jobs: linux: name: Haskell-CI - Linux - ${{ matrix.compiler }} @@ -44,19 +45,19 @@ jobs: strategy: matrix: include: - - compiler: ghc-9.8.0.20230822 + - compiler: ghc-9.8.1 compilerKind: ghc - compilerVersion: 9.8.0.20230822 + compilerVersion: 9.8.1 setup-method: ghcup - allow-failure: true - - compiler: ghc-9.6.2 + allow-failure: false + - compiler: ghc-9.6.3 compilerKind: ghc - compilerVersion: 9.6.2 + compilerVersion: 9.6.3 setup-method: ghcup allow-failure: false - - compiler: ghc-9.4.6 + - compiler: ghc-9.4.7 compilerKind: ghc - compilerVersion: 9.4.6 + compilerVersion: 9.4.7 setup-method: ghcup allow-failure: false - compiler: ghc-9.2.8 @@ -94,7 +95,6 @@ jobs: mkdir -p "$HOME/.ghcup/bin" curl -sL https://downloads.haskell.org/ghcup/0.1.19.5/x86_64-linux-ghcup-0.1.19.5 > "$HOME/.ghcup/bin/ghcup" chmod a+x "$HOME/.ghcup/bin/ghcup" - "$HOME/.ghcup/bin/ghcup" config add-release-channel https://raw.githubusercontent.com/haskell/ghcup-metadata/master/ghcup-prereleases-0.0.7.yaml; "$HOME/.ghcup/bin/ghcup" install ghc "$HCVER" || (cat "$HOME"/.ghcup/logs/*.* && false) "$HOME/.ghcup/bin/ghcup" install cabal 3.10.1.0 || (cat "$HOME"/.ghcup/logs/*.* && false) else @@ -104,7 +104,6 @@ jobs: mkdir -p "$HOME/.ghcup/bin" curl -sL https://downloads.haskell.org/ghcup/0.1.19.5/x86_64-linux-ghcup-0.1.19.5 > "$HOME/.ghcup/bin/ghcup" chmod a+x "$HOME/.ghcup/bin/ghcup" - "$HOME/.ghcup/bin/ghcup" config add-release-channel https://raw.githubusercontent.com/haskell/ghcup-metadata/master/ghcup-prereleases-0.0.7.yaml; "$HOME/.ghcup/bin/ghcup" install cabal 3.10.1.0 || (cat "$HOME"/.ghcup/logs/*.* && false) fi env: @@ -138,7 +137,7 @@ jobs: echo "HCNUMVER=$HCNUMVER" >> "$GITHUB_ENV" echo "ARG_TESTS=--enable-tests" >> "$GITHUB_ENV" echo "ARG_BENCH=--enable-benchmarks" >> "$GITHUB_ENV" - if [ $((HCNUMVER >= 90800)) -ne 0 ] ; then echo "HEADHACKAGE=true" >> "$GITHUB_ENV" ; else echo "HEADHACKAGE=false" >> "$GITHUB_ENV" ; fi + echo "HEADHACKAGE=false" >> "$GITHUB_ENV" echo "ARG_COMPILER=--$HCKIND --with-compiler=$HC" >> "$GITHUB_ENV" echo "GHCJSARITH=0" >> "$GITHUB_ENV" env: @@ -167,18 +166,6 @@ jobs: repository hackage.haskell.org url: http://hackage.haskell.org/ EOF - if $HEADHACKAGE; then - cat >> $CABAL_CONFIG <> $CABAL_CONFIG <> cabal.project cat >> cabal.project <> cabal.project - fi $HCPKG list --simple-output --names-only | perl -ne 'for (split /\s+/) { print "constraints: $_ installed\n" unless /^(agda-stdlib-utils)$/; }' >> cabal.project.local cat cabal.project cat cabal.project.local diff --git a/agda-stdlib-utils.cabal b/agda-stdlib-utils.cabal index a482bee069..3accc6972e 100644 --- a/agda-stdlib-utils.cabal +++ b/agda-stdlib-utils.cabal @@ -6,9 +6,9 @@ description: Helper programs for setting up the Agda standard library. license: MIT tested-with: - GHC == 9.8.0 - GHC == 9.6.2 - GHC == 9.4.6 + GHC == 9.8.1 + GHC == 9.6.3 + GHC == 9.4.7 GHC == 9.2.8 GHC == 9.0.2 GHC == 8.10.7 From 2fda34dde417144545aa72ffbfcf3328a616ecb1 Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Wed, 1 Nov 2023 02:32:30 +0000 Subject: [PATCH 44/57] [fixes #2175] Documentation misc. typos etc. for RC1 (#2183) * missing comma! * corrected reference to `README.Design.Decidability` * typo: capitalisation * updated `installation-guide` * added word to `NonZero` section heading * Run workflows on any PR * Add merge-group trigger to workflows --------- Co-authored-by: MatthewDaggitt --- CHANGELOG.md | 4 ++-- README.agda | 2 +- notes/installation-guide.md | 12 ++++++------ src/Relation/Nullary/Decidable/Core.agda | 4 ++-- 4 files changed, 11 insertions(+), 11 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 535c899aaf..a375b516ce 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -504,7 +504,7 @@ Non-backwards compatible changes and `Codata.Guarded.Stream.Relation.Binary.Pointwise`, the proofs called `lookup` have been renamed `lookup⁺`. -#### Changes to `Data.(Nat/Integer/Rational)` proofs of `NonZero`/`Positive`/`Negative` to instance arguments +#### Changes to `Data.(Nat/Integer/Rational)` proofs of `NonZero`/`Positive`/`Negative` to use 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. @@ -2766,7 +2766,7 @@ Additions to existing modules * Added new definitions and proofs to `Data.Nat.Primality`: ```agda Composite : ℕ → Set - composite? : Decidable composite + composite? : Decidable Composite composite⇒¬prime : Composite n → ¬ Prime n ¬composite⇒prime : 2 ≤ n → ¬ Composite n → Prime n prime⇒¬composite : Prime n → ¬ Composite n diff --git a/README.agda b/README.agda index cc436d0848..41b536a0d0 100644 --- a/README.agda +++ b/README.agda @@ -7,7 +7,7 @@ module README where -- -- Authors: Nils Anders Danielsson, Matthew Daggitt, Guillaume Allais -- with contributions from Andreas Abel, Stevan Andjelkovic, --- Jean-Philippe Bernardy, Peter Berry, Bradley Hardy Joachim Breitner, +-- Jean-Philippe Bernardy, Peter Berry, Bradley Hardy, Joachim Breitner, -- Samuel Bronson, Daniel Brown, Jacques Carette, James Chapman, -- Liang-Ting Chen, Dominique Devriese, Dan Doel, Érdi Gergő, -- Zack Grannan, Helmut Grohne, Simon Foster, Liyang Hu, Jason Hu, diff --git a/notes/installation-guide.md b/notes/installation-guide.md index f05f21ed90..67d1685ad9 100644 --- a/notes/installation-guide.md +++ b/notes/installation-guide.md @@ -3,19 +3,19 @@ Installation instructions Note: the full story on installing Agda libraries can be found at [readthedocs](http://agda.readthedocs.io/en/latest/tools/package-system.html). -Use version v1.7.2 of the standard library with Agda 2.6.3. +Use version v2.0 of the standard library with Agda 2.6.4. 1. Navigate to a suitable directory `$HERE` (replace appropriately) where you would like to install the library. -2. Download the tarball of v1.7.2 of the standard library. This can either be +2. Download the tarball of v2.0 of the standard library. This can either be done manually by visiting the Github repository for the library, or via the command line as follows: ``` - wget -O agda-stdlib.tar https://github.com/agda/agda-stdlib/archive/v1.7.2.tar.gz + wget -O agda-stdlib.tar https://github.com/agda/agda-stdlib/archive/v2.0.tar.gz ``` Note that you can replace `wget` with other popular tools such as `curl` and that - you can replace `1.7.2` with any other version of the library you desire. + you can replace `2.0` with any other version of the library you desire. 3. Extract the standard library from the tarball. Again this can either be done manually or via the command line as follows: @@ -26,7 +26,7 @@ Use version v1.7.2 of the standard library with Agda 2.6.3. 4. [ OPTIONAL ] If using [cabal](https://www.haskell.org/cabal/) then run the commands to install via cabal: ``` - cd agda-stdlib-1.7.2 + cd agda-stdlib-2.0 cabal install ``` @@ -40,7 +40,7 @@ Use version v1.7.2 of the standard library with Agda 2.6.3. 6. Register the standard library with Agda's package system by adding the following line to `$HOME/.agda/libraries`: ``` - $HERE/agda-stdlib-1.7.2/standard-library.agda-lib + $HERE/agda-stdlib-2.0/standard-library.agda-lib ``` Now, the standard library is ready to be used either: diff --git a/src/Relation/Nullary/Decidable/Core.agda b/src/Relation/Nullary/Decidable/Core.agda index 166b1f264e..ff73c6af20 100644 --- a/src/Relation/Nullary/Decidable/Core.agda +++ b/src/Relation/Nullary/Decidable/Core.agda @@ -35,8 +35,8 @@ private -- reflects the boolean result. This definition allows the boolean -- part of the decision procedure to compute independently from the -- proof. This leads to better computational behaviour when we only care --- about the result and not the proof. See README.Decidability for --- further details. +-- about the result and not the proof. See README.Design.Decidability +-- for further details. infix 2 _because_ From 188f74cc3f6e4edc36ee1ba816da8cb05b96178f Mon Sep 17 00:00:00 2001 From: jamesmckinna <31931406+jamesmckinna@users.noreply.github.com> Date: Wed, 1 Nov 2023 02:33:25 +0000 Subject: [PATCH 45/57] [fixes #2178] regularise and specify/systematise, the conventions for symbol usage (#2185) * regularise and specify/systematise, the conventions for symbol usage * typos/revisions --- CHANGELOG.md | 54 ++--- notes/style-guide.md | 44 +++- src/Algebra/Consequences/Propositional.agda | 54 ++--- src/Algebra/Consequences/Setoid.agda | 242 ++++++++++---------- 4 files changed, 216 insertions(+), 178 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index a375b516ce..557ec01b71 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -112,21 +112,21 @@ Non-backwards compatible changes 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` + - `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` + - `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` + - `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` + - `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. @@ -2157,16 +2157,16 @@ Additions to existing modules * 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 _•_ + 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 _•_ + 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 @@ -2176,15 +2176,15 @@ Additions to existing modules 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ˡ _◦_ + 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`: diff --git a/notes/style-guide.md b/notes/style-guide.md index d6201abe88..95ef7ba2d0 100644 --- a/notes/style-guide.md +++ b/notes/style-guide.md @@ -347,12 +347,12 @@ line of code, indented by two spaces. #### Variables -* `Level` and `Set`s can always be generalized using the keyword `variable`. +* `Level` and `Set`s can always be generalised using the keyword `variable`. * A file may only declare variables of other types if those types are used in the definition of the main type that the file concerns itself with. - At the moment the policy is *not* to generalize over any other types to - minimize the amount of information that users have to keep in their head + At the moment the policy is *not* to generalise over any other types to + minimise the amount of information that users have to keep in their head concurrently. * Example 1: the main type in `Data.List.Properties` is `List A` where `A : Set a`. @@ -443,6 +443,44 @@ word within a compound word is capitalized except for the first word. relations should be used over the `¬` symbol (e.g. `m+n≮n` should be used instead of `¬m+n Date: Wed, 1 Nov 2023 20:07:23 +0900 Subject: [PATCH 46/57] Move `T?` to `Relation.Nullary.Decidable.Core` (#2189) * Move `T?` to `Relation.Nullary.Decidable.Core` * Var name fix * Some refactoring * Fix failing tests and address remaining comments * Fix style-guide --- CHANGELOG.md | 9 ++++ README/Design/Hierarchies.agda | 4 +- notes/style-guide.md | 2 +- src/Data/Bool.agda | 3 -- src/Data/Bool/Base.agda | 18 ++++---- src/Data/Bool/Properties.agda | 43 ++++++++++--------- .../Relation/Binary/BagAndSetEquality.agda | 2 +- src/Effect/Applicative.agda | 2 +- src/Function/Bundles.agda | 2 +- src/Relation/Nullary.agda | 17 ++------ src/Relation/Nullary/Decidable/Core.agda | 5 ++- src/Relation/Nullary/Reflects.agda | 17 +++++--- 12 files changed, 67 insertions(+), 57 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 557ec01b71..46917d7117 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -838,6 +838,9 @@ Non-backwards compatible changes 4. The modules `Relation.Nullary.(Product/Sum/Implication)` have been deprecated and their contents moved to `Relation.Nullary.(Negation/Reflects/Decidable)`. + + 5. The proof `T?` has been moved from `Data.Bool.Properties` to `Relation.Nullary.Decidable.Core` + (but is still re-exported by the former). as well as the following breaking changes: @@ -3574,6 +3577,12 @@ Additions to existing modules poset : Set a → Poset _ _ _ ``` +* Added new proof in `Relation.Nullary.Reflects`: + ```agda + T-reflects : Reflects (T b) b + T-reflects-elim : Reflects (T a) b → b ≡ a + ``` + * Added new operations in `System.Exit`: ``` isSuccess : ExitCode → Bool diff --git a/README/Design/Hierarchies.agda b/README/Design/Hierarchies.agda index bc9220ba23..71ca5eed44 100644 --- a/README/Design/Hierarchies.agda +++ b/README/Design/Hierarchies.agda @@ -265,7 +265,7 @@ record Semigroup : Set (suc (a ⊔ ℓ)) where -- IsA A -- / || \ / || \ -- IsB IsC IsD B C D - + -- The procedure for re-exports in the bundles is as follows: -- 1. `open IsA isA public using (IsC, M)` where `M` is everything @@ -280,7 +280,7 @@ record Semigroup : Set (suc (a ⊔ ℓ)) where -- 5. `open B b public using (O)` where `O` is everything exported -- by `B` but not exported by `IsA`. - + -- 6. Construct `d : D` via the `isC` obtained in step 1. -- 7. `open D d public using (P)` where `P` is everything exported diff --git a/notes/style-guide.md b/notes/style-guide.md index 95ef7ba2d0..a3c918e6a4 100644 --- a/notes/style-guide.md +++ b/notes/style-guide.md @@ -402,7 +402,7 @@ word within a compound word is capitalized except for the first word. * Rational variables are named `p`, `q`, `r`, ... (default `p`) -* All other variables tend to be named `x`, `y`, `z`. +* All other variables should be named `x`, `y`, `z`. * Collections of elements are usually indicated by appending an `s` (e.g. if you are naming your variables `x` and `y` then lists diff --git a/src/Data/Bool.agda b/src/Data/Bool.agda index 1270125254..55ed146e2c 100644 --- a/src/Data/Bool.agda +++ b/src/Data/Bool.agda @@ -8,9 +8,6 @@ module Data.Bool where -open import Relation.Nullary -open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) - ------------------------------------------------------------------------ -- The boolean type and some operations diff --git a/src/Data/Bool/Base.agda b/src/Data/Bool/Base.agda index bc2e67644b..ea99ee880d 100644 --- a/src/Data/Bool/Base.agda +++ b/src/Data/Bool/Base.agda @@ -57,17 +57,19 @@ true xor b = not b false xor b = b ------------------------------------------------------------------------ --- Other operations - -infix 0 if_then_else_ - -if_then_else_ : Bool → A → A → A -if true then t else f = t -if false then t else f = f +-- Conversion to Set -- A function mapping true to an inhabited type and false to an empty -- type. - T : Bool → Set T true = ⊤ T false = ⊥ + +------------------------------------------------------------------------ +-- Other operations + +infix 0 if_then_else_ + +if_then_else_ : Bool → A → A → A +if true then t else f = t +if false then t else f = f diff --git a/src/Data/Bool/Properties.agda b/src/Data/Bool/Properties.agda index e911390366..bb8cda9a0e 100644 --- a/src/Data/Bool/Properties.agda +++ b/src/Data/Bool/Properties.agda @@ -28,8 +28,7 @@ open import Relation.Binary.Definitions using (Decidable; Reflexive; Transitive; Antisymmetric; Minimum; Maximum; Total; Irrelevant; Irreflexive; Asymmetric; Trans; Trichotomous; tri≈; tri<; tri>; _Respects₂_) open import Relation.Binary.PropositionalEquality.Core open import Relation.Binary.PropositionalEquality.Properties -open import Relation.Nullary.Reflects using (ofʸ; ofⁿ) -open import Relation.Nullary.Decidable.Core using (True; does; proof; yes; no) +open import Relation.Nullary.Decidable.Core using (True; yes; no; fromWitness) import Relation.Unary as U open import Algebra.Definitions {A = Bool} _≡_ @@ -726,15 +725,17 @@ xor-∧-commutativeRing = ⊕-∧-commutativeRing open XorRing _xor_ xor-is-ok ------------------------------------------------------------------------ --- Miscellaneous other properties +-- Properties of if_then_else_ -⇔→≡ : {x y z : Bool} → x ≡ z ⇔ y ≡ z → x ≡ y -⇔→≡ {true } {true } hyp = refl -⇔→≡ {true } {false} {true } hyp = sym (Equivalence.to hyp refl) -⇔→≡ {true } {false} {false} hyp = Equivalence.from hyp refl -⇔→≡ {false} {true } {true } hyp = Equivalence.from hyp refl -⇔→≡ {false} {true } {false} hyp = sym (Equivalence.to hyp refl) -⇔→≡ {false} {false} hyp = refl +if-float : ∀ (f : A → B) b {x y} → + f (if b then x else y) ≡ (if b then f x else f y) +if-float _ true = refl +if-float _ false = refl + +------------------------------------------------------------------------ +-- Properties of T + +open Relation.Nullary.Decidable.Core public using (T?) T-≡ : ∀ {x} → T x ⇔ x ≡ true T-≡ {false} = mk⇔ (λ ()) (λ ()) @@ -757,18 +758,20 @@ T-∨ {false} {false} = mk⇔ inj₁ [ id , id ] T-irrelevant : U.Irrelevant T T-irrelevant {true} _ _ = refl -T? : U.Decidable T -does (T? b) = b -proof (T? true ) = ofʸ _ -proof (T? false) = ofⁿ λ() - T?-diag : ∀ b → T b → True (T? b) -T?-diag true _ = _ +T?-diag b = fromWitness + +------------------------------------------------------------------------ +-- Miscellaneous other properties + +⇔→≡ : {x y z : Bool} → x ≡ z ⇔ y ≡ z → x ≡ y +⇔→≡ {true } {true } hyp = refl +⇔→≡ {true } {false} {true } hyp = sym (Equivalence.to hyp refl) +⇔→≡ {true } {false} {false} hyp = Equivalence.from hyp refl +⇔→≡ {false} {true } {true } hyp = Equivalence.from hyp refl +⇔→≡ {false} {true } {false} hyp = sym (Equivalence.to hyp refl) +⇔→≡ {false} {false} hyp = refl -if-float : ∀ (f : A → B) b {x y} → - f (if b then x else y) ≡ (if b then f x else f y) -if-float _ true = refl -if-float _ false = refl ------------------------------------------------------------------------ -- DEPRECATED NAMES diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 31ab6941ac..17af3d8955 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -44,7 +44,7 @@ import Relation.Binary.Reasoning.Preorder as PreorderReasoning open import Relation.Binary.PropositionalEquality as P using (_≡_; _≢_; _≗_; refl) open import Relation.Binary.Reasoning.Syntax -open import Relation.Nullary +open import Relation.Nullary.Negation using (¬_) open import Data.List.Membership.Propositional.Properties private diff --git a/src/Effect/Applicative.agda b/src/Effect/Applicative.agda index a2263c6e0a..95e2debda8 100644 --- a/src/Effect/Applicative.agda +++ b/src/Effect/Applicative.agda @@ -56,7 +56,7 @@ record RawApplicative (F : Set f → Set g) : Set (suc f ⊔ g) where -- Haskell-style alternative name for pure return : A → F A return = pure - + -- backwards compatibility: unicode variants _⊛_ : F (A → B) → F A → F B _⊛_ = _<*>_ diff --git a/src/Function/Bundles.agda b/src/Function/Bundles.agda index 9e18c30fe5..98aac7a48c 100644 --- a/src/Function/Bundles.agda +++ b/src/Function/Bundles.agda @@ -380,7 +380,7 @@ module _ (From : Setoid a ℓ₁) (To : Setoid b ℓ₂) where -- For further background on (split) surjections, one may consult any -- general mathematical references which work without the principle -- of choice. For example: - -- + -- -- https://ncatlab.org/nlab/show/split+epimorphism. -- -- The connection to set-theoretic notions with the same names is diff --git a/src/Relation/Nullary.agda b/src/Relation/Nullary.agda index b8a4c30470..0b0f83f77b 100644 --- a/src/Relation/Nullary.agda +++ b/src/Relation/Nullary.agda @@ -15,20 +15,9 @@ open import Agda.Builtin.Equality ------------------------------------------------------------------------ -- Re-exports -open import Relation.Nullary.Negation.Core public using - ( ¬_; _¬-⊎_ - ; contradiction; contradiction₂; contraposition - ) - -open import Relation.Nullary.Reflects public using - ( Reflects; ofʸ; ofⁿ - ; _×-reflects_; _⊎-reflects_; _→-reflects_ - ) - -open import Relation.Nullary.Decidable.Core public using - ( Dec; does; proof; yes; no; _because_; recompute - ; ¬?; _×-dec_; _⊎-dec_; _→-dec_ - ) +open import Relation.Nullary.Negation.Core public +open import Relation.Nullary.Reflects public +open import Relation.Nullary.Decidable.Core public ------------------------------------------------------------------------ -- Irrelevant types diff --git a/src/Relation/Nullary/Decidable/Core.agda b/src/Relation/Nullary/Decidable/Core.agda index ff73c6af20..f77613ea65 100644 --- a/src/Relation/Nullary/Decidable/Core.agda +++ b/src/Relation/Nullary/Decidable/Core.agda @@ -12,7 +12,7 @@ module Relation.Nullary.Decidable.Core where open import Level using (Level; Lift) -open import Data.Bool.Base using (Bool; false; true; not; T; _∧_; _∨_) +open import Data.Bool.Base using (Bool; T; false; true; not; _∧_; _∨_) open import Data.Unit.Base using (⊤) open import Data.Empty using (⊥) open import Data.Empty.Irrelevant using (⊥-elim) @@ -79,6 +79,9 @@ recompute (no ¬a) a = ⊥-elim (¬a a) infixr 1 _⊎-dec_ infixr 2 _×-dec_ _→-dec_ +T? : ∀ x → Dec (T x) +T? x = x because T-reflects x + ¬? : Dec A → Dec (¬ A) does (¬? a?) = not (does a?) proof (¬? a?) = ¬-reflects (proof a?) diff --git a/src/Relation/Nullary/Reflects.agda b/src/Relation/Nullary/Reflects.agda index 4f97bd8709..3b301fdbce 100644 --- a/src/Relation/Nullary/Reflects.agda +++ b/src/Relation/Nullary/Reflects.agda @@ -11,11 +11,12 @@ module Relation.Nullary.Reflects where open import Agda.Builtin.Equality open import Data.Bool.Base +open import Data.Unit.Base using (⊤) open import Data.Empty open import Data.Sum.Base using (_⊎_; inj₁; inj₂) open import Data.Product.Base using (_×_; _,_; proj₁; proj₂) open import Level using (Level) -open import Function.Base using (_$_; _∘_; const) +open import Function.Base using (_$_; _∘_; const; id) open import Relation.Nullary.Negation.Core @@ -54,28 +55,31 @@ invert (ofⁿ ¬a) = ¬a ------------------------------------------------------------------------ -- Interaction with negation, product, sums etc. +infixr 1 _⊎-reflects_ +infixr 2 _×-reflects_ _→-reflects_ + +T-reflects : ∀ b → Reflects (T b) b +T-reflects true = of _ +T-reflects false = of id + -- If we can decide A, then we can decide its negation. ¬-reflects : ∀ {b} → Reflects A b → Reflects (¬ A) (not b) ¬-reflects (ofʸ a) = ofⁿ (_$ a) ¬-reflects (ofⁿ ¬a) = ofʸ ¬a -- If we can decide A and Q then we can decide their product -infixr 2 _×-reflects_ _×-reflects_ : ∀ {a b} → Reflects A a → Reflects B b → Reflects (A × B) (a ∧ b) ofʸ a ×-reflects ofʸ b = ofʸ (a , b) ofʸ a ×-reflects ofⁿ ¬b = ofⁿ (¬b ∘ proj₂) ofⁿ ¬a ×-reflects _ = ofⁿ (¬a ∘ proj₁) - -infixr 1 _⊎-reflects_ _⊎-reflects_ : ∀ {a b} → Reflects A a → Reflects B b → Reflects (A ⊎ B) (a ∨ b) ofʸ a ⊎-reflects _ = ofʸ (inj₁ a) ofⁿ ¬a ⊎-reflects ofʸ b = ofʸ (inj₂ b) ofⁿ ¬a ⊎-reflects ofⁿ ¬b = ofⁿ (¬a ¬-⊎ ¬b) -infixr 2 _→-reflects_ _→-reflects_ : ∀ {a b} → Reflects A a → Reflects B b → Reflects (A → B) (not a ∨ b) ofʸ a →-reflects ofʸ b = ofʸ (const b) @@ -95,3 +99,6 @@ det (ofʸ a) (ofʸ _) = refl det (ofʸ a) (ofⁿ ¬a) = contradiction a ¬a det (ofⁿ ¬a) (ofʸ a) = contradiction a ¬a det (ofⁿ ¬a) (ofⁿ _) = refl + +T-reflects-elim : ∀ {a b} → Reflects (T a) b → b ≡ a +T-reflects-elim {a} r = det r (T-reflects a) From 9be7cea6abda2767d1047c265d9e6c621e843316 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Fri, 3 Nov 2023 12:43:34 +0900 Subject: [PATCH 47/57] Make decidable versions of sublist functions the default (#2186) * Make decdable versions of sublist functions the default * Remove imports Bool.Properties * Address comments --- src/Data/List/Base.agda | 209 ++++++++++-------- .../Unary/Unique/DecSetoid/Properties.agda | 4 +- src/Data/String/Base.agda | 18 +- src/Data/Vec/Base.agda | 16 +- 4 files changed, 133 insertions(+), 114 deletions(-) diff --git a/src/Data/List/Base.agda b/src/Data/List/Base.agda index 58b51e4893..16546110b6 100644 --- a/src/Data/List/Base.agda +++ b/src/Data/List/Base.agda @@ -23,11 +23,11 @@ open import Data.These.Base as These using (These; this; that; these) open import Function.Base using (id; _∘_ ; _∘′_; _∘₂_; _$_; const; flip) open import Level using (Level) -open import Relation.Nullary.Decidable.Core using (does; ¬?) open import Relation.Unary using (Pred; Decidable) open import Relation.Binary.Core using (Rel) import Relation.Binary.Definitions as B open import Relation.Binary.PropositionalEquality.Core using (_≡_) +open import Relation.Nullary.Decidable.Core using (T?; does; ¬?) private variable @@ -346,141 +346,160 @@ removeAt : (xs : List A) → Fin (length xs) → List A removeAt (x ∷ xs) zero = xs removeAt (x ∷ xs) (suc i) = x ∷ removeAt xs i --- The following are functions which split a list up using boolean --- predicates. However, in practice they are difficult to use and --- prove properties about, and are mainly provided for advanced use --- cases where one wants to minimise dependencies. In most cases --- you probably want to use the versions defined below based on --- decidable predicates. e.g. use `takeWhile (_≤? 10) xs` --- instead of `takeWhileᵇ (_≤ᵇ 10) xs` +------------------------------------------------------------------------ +-- Operations for filtering lists + +-- The following are a variety of functions that can be used to +-- construct sublists using a predicate. +-- +-- Each function has two forms. The first main variant uses a +-- proof-relevant decidable predicate, while the second variant uses +-- a irrelevant boolean predicate and are suffixed with a `ᵇ` character, +-- typed as \^b. +-- +-- The decidable versions have several advantages: 1) easier to prove +-- properties, 2) better meta-variable inference and 3) most of the rest +-- of the library is set-up to work with decidable predicates. However, +-- in rare cases the boolean versions can be useful, mainly when one +-- wants to minimise dependencies. +-- +-- In summary, in most cases you probably want to use the decidable +-- versions over the boolean versions, e.g. use `takeWhile (_≤? 10) xs` +-- rather than `takeWhileᵇ (_≤ᵇ 10) xs`. + +takeWhile : ∀ {P : Pred A p} → Decidable P → List A → List A +takeWhile P? [] = [] +takeWhile P? (x ∷ xs) with does (P? x) +... | true = x ∷ takeWhile P? xs +... | false = [] takeWhileᵇ : (A → Bool) → List A → List A -takeWhileᵇ p [] = [] -takeWhileᵇ p (x ∷ xs) = if p x then x ∷ takeWhileᵇ p xs else [] +takeWhileᵇ p = takeWhile (T? ∘ p) + +dropWhile : ∀ {P : Pred A p} → Decidable P → List A → List A +dropWhile P? [] = [] +dropWhile P? (x ∷ xs) with does (P? x) +... | true = dropWhile P? xs +... | false = x ∷ xs dropWhileᵇ : (A → Bool) → List A → List A -dropWhileᵇ p [] = [] -dropWhileᵇ p (x ∷ xs) = if p x then dropWhileᵇ p xs else x ∷ xs +dropWhileᵇ p = dropWhile (T? ∘ p) + +filter : ∀ {P : Pred A p} → Decidable P → List A → List A +filter P? [] = [] +filter P? (x ∷ xs) with does (P? x) +... | false = filter P? xs +... | true = x ∷ filter P? xs filterᵇ : (A → Bool) → List A → List A -filterᵇ p [] = [] -filterᵇ p (x ∷ xs) = if p x then x ∷ filterᵇ p xs else filterᵇ p xs +filterᵇ p = filter (T? ∘ p) + +partition : ∀ {P : Pred A p} → Decidable P → List A → (List A × List A) +partition P? [] = ([] , []) +partition P? (x ∷ xs) with does (P? x) | partition P? xs +... | true | (ys , zs) = (x ∷ ys , zs) +... | false | (ys , zs) = (ys , x ∷ zs) partitionᵇ : (A → Bool) → List A → List A × List A -partitionᵇ p [] = ([] , []) -partitionᵇ p (x ∷ xs) = (if p x then Prod.map₁ else Prod.map₂′) (x ∷_) (partitionᵇ p xs) +partitionᵇ p = partition (T? ∘ p) + +span : ∀ {P : Pred A p} → Decidable P → List A → (List A × List A) +span P? [] = ([] , []) +span P? ys@(x ∷ xs) with does (P? x) +... | true = Prod.map (x ∷_) id (span P? xs) +... | false = ([] , ys) + spanᵇ : (A → Bool) → List A → List A × List A -spanᵇ p [] = ([] , []) -spanᵇ p (x ∷ xs) = if p x - then Prod.map₁ (x ∷_) (spanᵇ p xs) - else ([] , x ∷ xs) +spanᵇ p = span (T? ∘ p) + +break : ∀ {P : Pred A p} → Decidable P → List A → (List A × List A) +break P? = span (¬? ∘ P?) breakᵇ : (A → Bool) → List A → List A × List A -breakᵇ p = spanᵇ (not ∘ p) +breakᵇ p = break (T? ∘ p) + +-- The predicate `P` represents the notion of newline character for the +-- type `A`. It is used to split the input list into a list of lines. +-- Some lines may be empty if the input contains at least two +-- consecutive newline characters. +linesBy : ∀ {P : Pred A p} → Decidable P → List A → List (List A) +linesBy {A = A} P? = go nothing where -linesByᵇ : (A → Bool) → List A → List (List A) -linesByᵇ {A = A} p = go nothing - where go : Maybe (List A) → List A → List (List A) go acc [] = maybe′ ([_] ∘′ reverse) [] acc - go acc (c ∷ cs) with p c + go acc (c ∷ cs) with does (P? c) ... | true = reverse (Maybe.fromMaybe [] acc) ∷ go nothing cs ... | false = go (just (c ∷ Maybe.fromMaybe [] acc)) cs -wordsByᵇ : (A → Bool) → List A → List (List A) -wordsByᵇ {A = A} p = go [] - where +linesByᵇ : (A → Bool) → List A → List (List A) +linesByᵇ p = linesBy (T? ∘ p) + +-- The predicate `P` represents the notion of space character for the +-- type `A`. It is used to split the input list into a list of words. +-- All the words are non empty and the output does not contain any space +-- characters. +wordsBy : ∀ {P : Pred A p} → Decidable P → List A → List (List A) +wordsBy {A = A} P? = go [] where + cons : List A → List (List A) → List (List A) cons [] ass = ass cons as ass = reverse as ∷ ass go : List A → List A → List (List A) go acc [] = cons acc [] - go acc (c ∷ cs) with p c + go acc (c ∷ cs) with does (P? c) ... | true = cons acc (go [] cs) ... | false = go (c ∷ acc) cs +wordsByᵇ : (A → Bool) → List A → List (List A) +wordsByᵇ p = wordsBy (T? ∘ p) + +derun : ∀ {R : Rel A p} → B.Decidable R → List A → List A +derun R? [] = [] +derun R? (x ∷ []) = x ∷ [] +derun R? (x ∷ xs@(y ∷ _)) with does (R? x y) | derun R? xs +... | true | ys = ys +... | false | ys = x ∷ ys + derunᵇ : (A → A → Bool) → List A → List A -derunᵇ r [] = [] -derunᵇ r (x ∷ []) = x ∷ [] -derunᵇ r (x ∷ y ∷ xs) = if r x y - then derunᵇ r (y ∷ xs) - else x ∷ derunᵇ r (y ∷ xs) +derunᵇ r = derun (T? ∘₂ r) + +deduplicate : ∀ {R : Rel A p} → B.Decidable R → List A → List A +deduplicate R? [] = [] +deduplicate R? (x ∷ xs) = x ∷ filter (¬? ∘ R? x) (deduplicate R? xs) deduplicateᵇ : (A → A → Bool) → List A → List A -deduplicateᵇ r [] = [] -deduplicateᵇ r (x ∷ xs) = x ∷ filterᵇ (not ∘ r x) (deduplicateᵇ r xs) +deduplicateᵇ r = deduplicate (T? ∘₂ r) -- Finds the first element satisfying the boolean predicate +find : ∀ {P : Pred A p} → Decidable P → List A → Maybe A +find P? [] = nothing +find P? (x ∷ xs) = if does (P? x) then just x else find P? xs + findᵇ : (A → Bool) → List A → Maybe A -findᵇ p [] = nothing -findᵇ p (x ∷ xs) = if p x then just x else findᵇ p xs +findᵇ p = find (T? ∘ p) -- Finds the index of the first element satisfying the boolean predicate -findIndexᵇ : (A → Bool) → (xs : List A) → Maybe $ Fin (length xs) -findIndexᵇ p [] = nothing -findIndexᵇ p (x ∷ xs) = if p x +findIndex : ∀ {P : Pred A p} → Decidable P → (xs : List A) → Maybe $ Fin (length xs) +findIndex P? [] = nothing +findIndex P? (x ∷ xs) = if does (P? x) then just zero - else Maybe.map suc (findIndexᵇ p xs) + else Maybe.map suc (findIndex P? xs) + +findIndexᵇ : (A → Bool) → (xs : List A) → Maybe $ Fin (length xs) +findIndexᵇ p = findIndex (T? ∘ p) -- Finds indices of all the elements satisfying the boolean predicate -findIndicesᵇ : (A → Bool) → (xs : List A) → List $ Fin (length xs) -findIndicesᵇ p [] = [] -findIndicesᵇ p (x ∷ xs) = if p x +findIndices : ∀ {P : Pred A p} → Decidable P → (xs : List A) → List $ Fin (length xs) +findIndices P? [] = [] +findIndices P? (x ∷ xs) = if does (P? x) then zero ∷ indices else indices - where indices = map suc (findIndicesᵇ p xs) - --- Equivalent functions that use a decidable predicate instead of a --- boolean function. - -takeWhile : ∀ {P : Pred A p} → Decidable P → List A → List A -takeWhile P? = takeWhileᵇ (does ∘ P?) - -dropWhile : ∀ {P : Pred A p} → Decidable P → List A → List A -dropWhile P? = dropWhileᵇ (does ∘ P?) + where indices = map suc (findIndices P? xs) -filter : ∀ {P : Pred A p} → Decidable P → List A → List A -filter P? = filterᵇ (does ∘ P?) - -partition : ∀ {P : Pred A p} → Decidable P → List A → (List A × List A) -partition P? = partitionᵇ (does ∘ P?) - -span : ∀ {P : Pred A p} → Decidable P → List A → (List A × List A) -span P? = spanᵇ (does ∘ P?) - -break : ∀ {P : Pred A p} → Decidable P → List A → (List A × List A) -break P? = breakᵇ (does ∘ P?) - --- The predicate `P` represents the notion of newline character for the --- type `A`. It is used to split the input list into a list of lines. --- Some lines may be empty if the input contains at least two --- consecutive newline characters. -linesBy : ∀ {P : Pred A p} → Decidable P → List A → List (List A) -linesBy P? = linesByᵇ (does ∘ P?) - --- The predicate `P` represents the notion of space character for the --- type `A`. It is used to split the input list into a list of words. --- All the words are non empty and the output does not contain any space --- characters. -wordsBy : ∀ {P : Pred A p} → Decidable P → List A → List (List A) -wordsBy P? = wordsByᵇ (does ∘ P?) - -derun : ∀ {R : Rel A p} → B.Decidable R → List A → List A -derun R? = derunᵇ (does ∘₂ R?) - -deduplicate : ∀ {R : Rel A p} → B.Decidable R → List A → List A -deduplicate R? = deduplicateᵇ (does ∘₂ R?) - -find : ∀ {P : Pred A p} → Decidable P → List A → Maybe A -find P? = findᵇ (does ∘ P?) - -findIndex : ∀ {P : Pred A p} → Decidable P → (xs : List A) → Maybe $ Fin (length xs) -findIndex P? = findIndexᵇ (does ∘ P?) - -findIndices : ∀ {P : Pred A p} → Decidable P → (xs : List A) → List $ Fin (length xs) -findIndices P? = findIndicesᵇ (does ∘ P?) +findIndicesᵇ : (A → Bool) → (xs : List A) → List $ Fin (length xs) +findIndicesᵇ p = findIndices (T? ∘ p) ------------------------------------------------------------------------ -- Actions on single elements diff --git a/src/Data/List/Relation/Unary/Unique/DecSetoid/Properties.agda b/src/Data/List/Relation/Unary/Unique/DecSetoid/Properties.agda index 9e97527ff8..73d42b7b8a 100644 --- a/src/Data/List/Relation/Unary/Unique/DecSetoid/Properties.agda +++ b/src/Data/List/Relation/Unary/Unique/DecSetoid/Properties.agda @@ -12,7 +12,6 @@ open import Data.List.Relation.Unary.All.Properties using (all-filter) open import Data.List.Relation.Unary.Unique.Setoid.Properties open import Level open import Relation.Binary.Bundles using (DecSetoid) -open import Relation.Nullary.Decidable using (¬?) module Data.List.Relation.Unary.Unique.DecSetoid.Properties where @@ -30,5 +29,4 @@ module _ (DS : DecSetoid a ℓ) where deduplicate-! : ∀ xs → Unique (deduplicate _≟_ xs) deduplicate-! [] = [] - deduplicate-! (x ∷ xs) = all-filter (λ y → ¬? (x ≟ y)) (deduplicate _≟_ xs) - ∷ filter⁺ S (λ y → ¬? (x ≟ y)) (deduplicate-! xs) + deduplicate-! (x ∷ xs) = all-filter _ (deduplicate _≟_ xs) ∷ filter⁺ S _ (deduplicate-! xs) diff --git a/src/Data/String/Base.agda b/src/Data/String/Base.agda index 1a56376f59..d583da042c 100644 --- a/src/Data/String/Base.agda +++ b/src/Data/String/Base.agda @@ -22,7 +22,7 @@ open import Level using (Level; 0ℓ) open import Relation.Binary.Core using (Rel) open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) open import Relation.Unary using (Pred; Decidable) -open import Relation.Nullary.Decidable.Core using (does) +open import Relation.Nullary.Decidable.Core using (does; T?) ------------------------------------------------------------------------ -- From Agda.Builtin: type and renamed primitives @@ -160,11 +160,11 @@ fromAlignment Right = padLeft ' ' ------------------------------------------------------------------------ -- Splitting strings -wordsByᵇ : (Char → Bool) → String → List String -wordsByᵇ p = List.map fromList ∘ List.wordsByᵇ p ∘ toList - wordsBy : ∀ {p} {P : Pred Char p} → Decidable P → String → List String -wordsBy P? = wordsByᵇ (does ∘ P?) +wordsBy P? = List.map fromList ∘ List.wordsBy P? ∘ toList + +wordsByᵇ : (Char → Bool) → String → List String +wordsByᵇ p = wordsBy (T? ∘ p) words : String → List String words = wordsByᵇ Char.isSpace @@ -173,11 +173,11 @@ words = wordsByᵇ Char.isSpace _ : words " abc b " ≡ "abc" ∷ "b" ∷ [] _ = refl -linesByᵇ : (Char → Bool) → String → List String -linesByᵇ p = List.map fromList ∘ List.linesByᵇ p ∘ toList - linesBy : ∀ {p} {P : Pred Char p} → Decidable P → String → List String -linesBy P? = linesByᵇ (does ∘ P?) +linesBy P? = List.map fromList ∘ List.linesBy P? ∘ toList + +linesByᵇ : (Char → Bool) → String → List String +linesByᵇ p = linesBy (T? ∘ p) lines : String → List String lines = linesByᵇ ('\n' Char.≈ᵇ_) diff --git a/src/Data/Vec/Base.agda b/src/Data/Vec/Base.agda index 1ecf91b9e7..10684dee30 100644 --- a/src/Data/Vec/Base.agda +++ b/src/Data/Vec/Base.agda @@ -8,7 +8,7 @@ module Data.Vec.Base where -open import Data.Bool.Base using (Bool; if_then_else_) +open import Data.Bool.Base using (Bool; true; false; if_then_else_) open import Data.Nat.Base open import Data.Fin.Base using (Fin; zero; suc) open import Data.List.Base as List using (List) @@ -17,7 +17,7 @@ open import Data.These.Base as These using (These; this; that; these) open import Function.Base using (const; _∘′_; id; _∘_) open import Level using (Level) open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; trans; cong) -open import Relation.Nullary.Decidable.Core using (does) +open import Relation.Nullary.Decidable.Core using (does; T?) open import Relation.Unary using (Pred; Decidable) private @@ -230,12 +230,14 @@ foldl₁ _⊕_ (x ∷ xs) = foldl _ _⊕_ x xs sum : Vec ℕ n → ℕ sum = foldr _ _+_ 0 -countᵇ : (A → Bool) → Vec A n → ℕ -countᵇ p [] = zero -countᵇ p (x ∷ xs) = if p x then suc (countᵇ p xs) else countᵇ p xs - count : ∀ {P : Pred A p} → Decidable P → Vec A n → ℕ -count P? = countᵇ (does ∘ P?) +count P? [] = zero +count P? (x ∷ xs) with does (P? x) +... | true = suc (count P? xs) +... | false = count P? xs + +countᵇ : (A → Bool) → Vec A n → ℕ +countᵇ p = count (T? ∘ p) ------------------------------------------------------------------------ -- Operations for building vectors From a45f04daaebbbbe57bce7f22ed2c3fa68f01cf27 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Thu, 14 Dec 2023 04:03:17 +0000 Subject: [PATCH 48/57] new `CHANGELOG` --- CHANGELOG.md | 3801 +------------------------------------------------- 1 file changed, 14 insertions(+), 3787 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 30487a5e19..3ddcb555f6 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1,3807 +1,34 @@ -Version 2.0-rc1 +Version 2.1-dev =============== -The library has been tested using Agda 2.6.4. - -NOTE: Version `2.0` contains various breaking changes and is not backwards -compatible with code written with version `1.X` of the library. +The library has been tested using Agda 2.6.4 and 2.6.4.1. Highlights ---------- -* 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. - -* 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. - -* Significant tightening of `import` statements internally in the library, - drastically decreasing the dependencies and hence load time of many key - modules. - -* 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. - Bug-fixes --------- -* In `Algebra.Definitions.RawSemiring` the record `Prime` did not - enforce that the number was not divisible by `1#`. To fix this - `p∤1 : p ∤ 1#` has been added as a field. - -* 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.Fin.Properties` the `i` argument to `opposite-suc` was implicit - but could not be inferred in general. It has been made explicit. - -* In `Data.List.Membership.Setoid` the operations `_∷=_` and `_─_` - had an extraneous `{A : Set a}` parameter. This has been removed. - -* In `Data.List.Relation.Ternary.Appending.Setoid` the constructors - were re-exported in their full generality which lead to unsolved meta - variables at their use sites. Now versions of the constructors specialised - to use the setoid's carrier set are re-exported. - -* In `Data.Nat.DivMod` the parameter `o` in the proof `/-monoˡ-≤` was - implicit but not inferrable. It has been changed to be explicit. - -* In `Data.Nat.DivMod` the parameter `m` in the proof `+-distrib-/-∣ʳ` was - implicit but not inferrable, while `n` is explicit but inferrable. - They have been to explicit and implicit respectively. - -* In `Data.Nat.GeneralisedArithmetic` the `s` and `z` arguments to the - following functions were implicit but not inferrable: - `fold-+`, `fold-k`, `fold-*`, `fold-pull`. They have been made explicit. - -* In `Data.Rational(.Unnormalised).Properties` the module `≤-Reasoning` - exported equality combinators using the generic setoid symbol `_≈_`. They - have been renamed to use the same `_≃_` symbol used for non-propositional - equality over `Rational`s, i.e. - ```agda - step-≈ ↦ step-≃ - step-≈˘ ↦ step-≃˘ - ``` - with corresponding associated syntax: - ```agda - _≈⟨_⟩_ ↦ _≃⟨_⟩_ - _≈⟨_⟨_ ↦ _≃⟨_⟨_ - ``` - -* In `Function.Construct.Composition` the combinators - `_⟶-∘_`, `_↣-∘_`, `_↠-∘_`, `_⤖-∘_`, `_⇔-∘_`, `_↩-∘_`, `_↪-∘_`, `_↔-∘_` - had their arguments in the wrong order. They have been flipped so they can - actually be used as a composition operator. - -* 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. - -* 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. - Non-backwards compatible changes -------------------------------- -### Removed deprecated names - -* All modules and names that were deprecated in `v1.2` and before have - been removed. - -### Changes to `LeftCancellative` and `RightCancellative` in `Algebra.Definitions` - -* 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` - -### Changes to ring structures in `Algebra` - -* 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*` - -### Refactoring of lattices in `Algebra.Structures/Bundles` 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. - -* The definition of `IsDistributiveLattice` and `IsBooleanAlgebra` have changed so - that they are no longer right-biased which hindered compositionality. - More concretely, `IsDistributiveLattice` now has fields: - ```agda - ∨-distrib-∧ : ∨ DistributesOver ∧ - ∧-distrib-∨ : ∧ DistributesOver ∨ - ``` - instead of - ```agda - ∨-distribʳ-∧ : ∨ DistributesOverʳ ∧ - ``` - and `IsBooleanAlgebra` now has fields: - ```agda - ∨-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. - -* Finally, the following auxiliary files have been moved: - ```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 - ``` - -#### Changes to `Algebra.Morphism.Structures` - -* Previously the record definitions: - ``` - IsNearSemiringHomomorphism - IsSemiringHomomorphism - IsRingHomomorphism - ``` - all had two separate proofs of `IsRelHomomorphism` within them. - -* To fix this they have all been redefined to build up from `IsMonoidHomomorphism`, - `IsNearSemiringHomomorphism`, and `IsSemiringHomomorphism` respectively, - adding a single property at each step. - -* Similarly, `IsLatticeHomomorphism` is now built as - `IsRelHomomorphism` along with proofs that `_∧_` and `_∨_` are homomorphic. - -* Finally `⁻¹-homo` in `IsRingHomomorphism` has been renamed to `-‿homo`. - -#### Renamed `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. - -* Full list of moved modules: - ```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 - ``` - -* Full list of 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 - ``` - -### Refactoring of the unindexed Functor/Applicative/Monad hierarchy in `Effect` - -* 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` + `_>>=_`. - 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. - - -#### 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 `Codata.Sized`, e.g. `Codata.Stream` - has become `Codata.Sized.Stream`. - -### New proof-irrelevant for empty type in `Data.Empty` - -* 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 the following two definitions have been modified: - - + 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 - ``` - -### Deprecation of `_≺_` in `Data.Fin.Base` - -* In `Data.Fin.Base` the relation `_≺_` and its single constructor `_≻toℕ_` - have been deprecated in favour of their extensional equivalent `_<_` - but omitting the inversion principle which pattern matching on `_≻toℕ_` - would achieve; this instead is proxied by the property `Data.Fin.Properties.toℕ<`. - -* Consequently in `Data.Fin.Induction`: - ``` - ≺-Rec - ≺-wellFounded - ≺-recBuilder - ≺-rec - ``` - these functions are also deprecated. - -* Likewise in `Data.Fin.Properties` the proofs `≺⇒<′` and `<′⇒≺` have been deprecated - in favour of their proxy counterparts `<⇒<′` and `<′⇒<`. - -### Standardisation of `insertAt`/`updateAt`/`removeAt` in `Data.List`/`Data.Vec` - -* Previously, the names and argument order of index-based insertion, update and removal functions for - various types of lists and vectors were inconsistent. - -* To fix this the names have all been standardised to `insertAt`/`updateAt`/`removeAt`. - -* Correspondingly the following changes have occurred: - -* In `Data.List.Base` the following have been added: - ```agda - insertAt : (xs : List A) → Fin (suc (length xs)) → A → List A - updateAt : (xs : List A) → Fin (length xs) → (A → A) → List A - removeAt : (xs : List A) → Fin (length xs) → List A - ``` - and the following has been deprecated - ``` - _─_ ↦ removeAt - ``` - -* In `Data.Vec.Base`: - ```agda - insert ↦ insertAt - remove ↦ removeAt - - updateAt : Fin n → (A → A) → Vec A n → Vec A n - ↦ - updateAt : Vec A n → Fin n → (A → A) → Vec A n - ``` - -* In `Data.Vec.Functional`: - ```agda - remove : Fin (suc n) → Vector A (suc n) → Vector A n - ↦ - removeAt : Vector A (suc n) → Fin (suc n) → Vector A n - - updateAt : Fin n → (A → A) → Vector A n → Vector A n - ↦ - updateAt : Vector A n → Fin n → (A → A) → Vector A n - ``` - -* The old names (and the names of all proofs about these functions) have been deprecated appropriately. - -#### Standardisation of `lookup` in `Data.(List/Vec/...)` - -* All the types of `lookup` functions (and variants) in the following modules - have been changed to match the argument convention adopted in the `List` module (i.e. - `lookup` takes its container first and the index, whose type may depend on the - container value, second): - ``` - 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 - ``` - -* To accommodate this in in `Data.Vec.Relation.Unary.Linked.Properties` - and `Codata.Guarded.Stream.Relation.Binary.Pointwise`, the proofs - called `lookup` have been renamed `lookup⁺`. - -#### Changes to `Data.(Nat/Integer/Rational)` proofs of `NonZero`/`Positive`/`Negative` to use 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 discussion for a fuller explanation of the motivation and implementation. - -* For example, whereas the type of division over `ℕ` used to be: - ```agda - _/_ : (dividend divisor : ℕ) {≢0 : False (divisor ≟ 0)} → ℕ - ``` - it is now: - ```agda - _/_ : (dividend divisor : ℕ) .{{_ : NonZero divisor}} → ℕ - ``` - -* This 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. - ```agda - 0/n≡0 : 0 / suc n ≡ 0 - ``` - - 2. *Add an instance argument parameter* - 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. - ```agda - 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. - ```agda - 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 form `suc n`. - 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, in `Data.Rational.Unnormalised.Base` the definition of `_≢0` is now - redundant and so has been removed. Additionally the following proofs about it have - also 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ℚᵘ - ``` - -### Changes to the definition of `_≤″_` in `Data.Nat.Base` (issue #1919) - -* The definition of `_≤″_` was previously: - ```agda - record _≤″_ (m n : ℕ) : Set where - constructor less-than-or-equal - field - {k} : ℕ - proof : m + k ≡ n - ``` - which introduced a spurious additional definition, when this is in fact, modulo - field names and implicit/explicit qualifiers, equivalent to the definition of left- - divisibility, `_∣ˡ_` for the `RawMagma` structure of `_+_`. - -* Since the addition of raw bundles to `Data.X.Base`, this definition can now be - made directly. Accordingly, the definition has been changed to: - ```agda - _≤″_ : (m n : ℕ) → Set - _≤″_ = _∣ˡ_ +-rawMagma - - pattern less-than-or-equal {k} prf = k , prf - ``` - -* Knock-on consequences include the need to retain the old constructor - name, now introduced as a pattern synonym, and introduction of (a function - equivalent to) the former field name/projection function `proof` as - `≤″-proof` in `Data.Nat.Properties`. - -### Changes to definition of `Prime` in `Data.Nat.Primality` - -* The definition of `Prime` 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 - ``` - -### Changes to operation reduction behaviour in `Data.Rational(.Unnormalised)` - -* 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 either on this reduction behaviour - or on eta-equality may no longer type-check. - -* 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. - -* Similarly, in order to prevent reduction, the equality `_≃_` in `Data.Rational.Base` - has been made into a data type with the single constructor `*≡*`. The destructor - `drop-*≡*` has been added to `Data.Rational.Properties`. - -#### Deprecation of old `Function` hierarchy - -* The new `Function` hierarchy was introduced in `v1.2` which follows - the same `Core`/`Definitions`/`Bundles`/`Structures` as all the other - hierarchies in the library. - -* At the time the old function hierarchy in: - ``` - Function.Equality - Function.Injection - Function.Surjection - Function.Bijection - Function.LeftInverse - Function.Inverse - Function.HalfAdjointEquivalence - Function.Related - ``` - was unofficially deprecated, but could not be officially deprecated because - of it's widespread use in the rest of the library. - -* Now, the old hierarchy modules have all been officially deprecated. All - uses of them in the rest of the library have been switched to use the - new hierarchy. - -* The latter is unfortunately a relatively big breaking change, but was judged - to be unavoidable given how widely used the old hierarchy was. - -#### Changes to the new `Function` hierarchy - -* In `Function.Bundles` the names of the fields in the records have been - changed from `f`, `g`, `cong₁` and `cong₂` to `to`, `from`, `to-cong`, `from-cong`. - -* In `Function.Definitions`, the module 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 two modules have - been removed and their definitions folded into `Function.Definitions`. - -* In `Function.Definitions` the following definitions 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)`. - -### New `Function.Strict` - -* 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`. - -### Change to the definition of `WfRec` in `Induction.WellFounded` (issue #2083) - -* Previously, the definition of `WfRec` was: - ```agda - WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _ - WfRec _<_ P x = ∀ y → y < x → P y - ``` - which meant that all arguments involving accessibility and wellfoundedness proofs - were polluted by almost-always-inferrable explicit arguments for the `y` position. - -* The definition has now been changed to make that argument *implicit*, as - ```agda - WfRec : Rel A r → ∀ {ℓ} → RecStruct A ℓ _ - WfRec _<_ P x = ∀ {y} → y < x → P y - ``` - -### Reorganisation 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. - - -### 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 about these definitions were spread over - `Relation.Nullary.(Negation/Product/Sum/Implication etc.)`. - -* To fix this all the contents of the latter is now exported by `Relation.Nullary`. - -* In order to achieve this the following backwards compatible changes have been made: - - 1. the definition of `Dec` and `recompute` have been moved to `Relation.Nullary.Decidable.Core` - - 2. the definition of `Reflects` has been moved to `Relation.Nullary.Reflects` - - 3. the definition of `¬_` has been moved to `Relation.Nullary.Negation.Core` - - 4. The modules `Relation.Nullary.(Product/Sum/Implication)` have been deprecated - and their contents moved to `Relation.Nullary.(Negation/Reflects/Decidable)`. - - 5. The proof `T?` has been moved from `Data.Bool.Properties` to `Relation.Nullary.Decidable.Core` - (but is still re-exported by the former). - - as well as the following breaking changes: - - 1. `¬?` has been moved from `Relation.Nullary.Negation.Core` to - `Relation.Nullary.Decidable.Core` - - 2. `¬-reflects` has been moved from `Relation.Nullary.Negation.Core` to - `Relation.Nullary.Reflects`. - - 3. `decidable-stable`, `excluded-middle` and `¬-drop-Dec` have been moved - from `Relation.Nullary.Negation` to `Relation.Nullary.Decidable`. - - 4. `fromDec` and `toDec` have been moved from `Data.Sum.Base` to `Data.Sum`. - - -### (Issue #2096) Introduction of flipped and negated relation symbols to bundles in `Relation.Binary.Bundles` - -* Previously, bundles such as `Preorder`, `Poset`, `TotalOrder` etc. did not have the flipped - and negated versions of the operators exposed. In some cases they could obtained by opening the - relevant `Relation.Binary.Properties.X` file but usually they had to be redefined every time. - -* To fix this, these bundles now all export all 4 versions of the operator: normal, converse, negated, - converse-negated. Accordingly they are no longer exported from the corresponding `Properties` file. - -* To make this work for `Preorder`, it was necessary to change the name of the relation symbol. - Previously, the symbol was `_∼_` which is (notationally) symmetric, so that its - converse relation could only be discussed *semantically* in terms of `flip _∼_`. - -* Now, the `Preorder` record field `_∼_` has been renamed to `_≲_`, with `_≳_` - introduced as a definition in `Relation.Binary.Bundles.Preorder`. - Partial backwards compatible has been achieved by redeclaring a deprecated version - of the old symbol in the record. Therefore, only _declarations_ of `PartialOrder` records will - need their field names updating. - - -### Changes to definition of `IsStrictTotalOrder` in `Relation.Binary.Structures` - -* The previous definition of the record `IsStrictTotalOrder` did not - build upon `IsStrictPartialOrder` as would be expected. - Instead it omitted several fields like irreflexivity as they were derivable from the - proof of trichotomy. However, this led to problems further up the hierarchy where - bundles such as `StrictTotalOrder` which contained multiple distinct proofs of - `IsStrictPartialOrder`. - -* To remedy this the definition of `IsStrictTotalOrder` has been changed to so - that it builds upon `IsStrictPartialOrder` as would be expected. - -* To aid migration, the old record definition has been moved to - `Relation.Binary.Structures.Biased` which contains the `isStrictTotalOrderᶜ` - smart constructor (which is re-exported by `Relation.Binary`) . Therefore the old code: - ```agda - <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ - <-isStrictTotalOrder = record - { isEquivalence = isEquivalence - ; trans = <-trans - ; compare = <-cmp - } - ``` - can be migrated either by updating to the new record fields if you have a proof of `IsStrictPartialOrder` - available: - ```agda - <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ - <-isStrictTotalOrder = record - { isStrictPartialOrder = <-isStrictPartialOrder - ; compare = <-cmp - } - ``` - or simply applying the smart constructor to the record definition as follows: - ```agda - <-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_ - <-isStrictTotalOrder = isStrictTotalOrderᶜ record - { isEquivalence = isEquivalence - ; trans = <-trans - ; compare = <-cmp - } - ``` - -### Changes to the interface of `Relation.Binary.Reasoning.Triple` - -* The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof - that the strict relation is irreflexive. - -* This allows the addition of 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 to derive anything. - -* Specialised versions of this combinator are available in the `Reasoning` modules - exported by the following 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 - ``` - -### A more modular design for `Relation.Binary.Reasoning` - -* Previously, every `Reasoning` module in the library tended to roll it's own set - of syntax for the combinators. This hindered consistency and maintainability. - -* To improve the situation, a new module `Relation.Binary.Reasoning.Syntax` - has been introduced which exports a wide range of sub-modules containing - pre-existing reasoning combinator syntax. - -* This makes it possible to add new or rename existing reasoning combinators to a - pre-existing `Reasoning` module in just a couple of lines - (e.g. see `∣-Reasoning` in `Data.Nat.Divisibility`) - -* One pre-requisite for that is that `≡-Reasoning` has been moved from - `Relation.Binary.PropositionalEquality.Core` (which shouldn't be - imported anyway as it's a `Core` module) to - `Relation.Binary.PropositionalEquality.Properties`. - It is still exported by `Relation.Binary.PropositionalEquality`. - -### Renaming of symmetric combinators in `Reasoning` modules - -* We've had various complaints about the symmetric version of reasoning combinators - that use the syntax `_R˘⟨_⟩_` for some relation `R`, (e.g. `_≡˘⟨_⟩_` and `_≃˘⟨_⟩_`) - introduced in `v1.0`. In particular: - 1. The symbol `˘` is hard to type. - 2. The symbol `˘` doesn't convey the direction of the equality - 3. The left brackets aren't vertically aligned with the left brackets of the non-symmetric version. - -* To address these problems we have renamed all the symmetric versions of the - combinators from `_R˘⟨_⟩_` to `_R⟨_⟨_` (the direction of the last bracket is flipped - to indicate the quality goes from right to left). - -* The old combinators still exist but have been deprecated. However due to - [Agda issue #5617](https://github.com/agda/agda/issues/5617), the deprecation warnings - don't fire correctly. We will not remove the old combinators before the above issue is - addressed. However, we still encourage migrating to the new combinators! +Other major improvements +------------------------ -* On a Linux-based system, the following command was used to globally migrate all uses of the - old combinators to the new ones in the standard library itself. - It *may* be useful when trying to migrate your own code: - ```bash - find . -type f -name "*.agda" -print0 | xargs -0 sed -i -e 's/˘⟨\(.*\)⟩/⟨\1⟨/g' - ``` - USUAL CAVEATS: It has not been widely tested and the standard library developers are not - responsible for the results of this command. It is strongly recommended you back up your - work before you attempt to run it. +Deprecated modules +------------------ -* NOTE: this refactoring may require some extra bracketing around the operator `_⟨_⟩_` from - `Function.Base` if the `_⟨_⟩_` operator is used within the reasoning proof. The symptom - for this is a `Don't know how to parse` error. +Deprecated names +---------------- -### Improvements to `Text.Pretty` and `Text.Regex` +New modules +----------- -* 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: +* The 'no infinite descent' principle for (accessible elements of) well-founded relations: ``` - 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 + Induction.NoInfiniteDescent ``` -* 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". - -### 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. - -* In `Algebra.Core` the operations `Opₗ` and `Opᵣ` have moved to `Algebra.Module.Core`. - -* 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. - -* In `Data.Fin.Base` the relations `_≤_` `_≥_` `_<_` `_>_` 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. - -* In `Data.Fin.Properties` the proof `inj⇒≟` that an injection from a type - `A` into `Fin n` induces a decision procedure for `_≡_` on `A` has been - generalised to other equivalences over `A` (i.e. to arbitrary setoids), and - renamed from `eq?` to the more descriptive and `inj⇒decSetoid`. - -* In `Data.Fin.Properties` the proof `pigeonhole` has been strengthened - so that the a proof that `i < j` rather than a mere `i ≢ j` is returned. - -* In `Data.Fin.Substitution.TermSubst` the fixity of `_/Var_` has been changed - from `infix 8` to `infixl 8`. - -* In `Data.Integer.DivMod` the previous implementations of - `_divℕ_`, `_div_`, `_modℕ_`, `_mod_` 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. +Additions to existing modules +----------------------------- -* In `Data.Integer.Properties` the first two arguments of `m≡n⇒m-n≡0` - (now renamed `i≡j⇒i-j≡0`) have been made implicit. - -* In `Data.(List|Vec).Relation.Binary.Lex.Strict` the argument `xs` in - `xs≮[]` in introduced in PRs #1648 and #1672 has now been made implicit. - -* In `Data.List.NonEmpty` the functions `split`, `flatten` and `flatten-split` - have been removed from. 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. - -* In `Data.Nat.Divisibility` the proof `m/n/o≡m/[n*o]` has been removed. In it's - place a new more general proof `m/n/o≡m/[n*o]` has been added to `Data.Nat.DivMod` - that doesn't require the `n * o ∣ m` pre-condition. - -* In `Data.Product.Relation.Binary.Lex.Strict` the proof of wellfoundedness - of the lexicographic ordering on products, no longer - requires the assumption of symmetry for the first equality relation `_≈₁_`. - -* In `Data.Rational.Base` the constructors `+0` and `+[1+_]` from `Data.Integer.Base` - are no longer re-exported by `Data.Rational.Base`. You will have to open `Data.Integer(.Base)` - directly to use them. - -* In `Data.Rational(.Unnormalised).Properties` the types of the proofs - `pos⇒1/pos`/`1/pos⇒pos` and `neg⇒1/neg`/`1/neg⇒neg` 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: - ```agda - pos⇒1/pos : ∀ p .{{_ : NonZero p}} .{{Positive (1/ p)}} → Positive p - 1/pos⇒pos : ∀ p .{{_ : Positive p}} → Positive (1/ p) - ``` - but are now: - ```agda - pos⇒1/pos : ∀ p .{{_ : Positive p}} → Positive (1/ p) - 1/pos⇒pos : ∀ p .{{_ : NonZero p}} .{{Positive (1/ p)}} → Positive p - ``` - -* In `Data.Sum.Base` the definitions `fromDec` and `toDec` have been moved to `Data.Sum`. - -* In `Data.Vec.Base` the definition of `_>>=_` under `Data.Vec.Base` has been - moved to the submodule `CartesianBind` in order to avoid clashing with the - new, alternative definition of `_>>=_`, located in the second new submodule - `DiagonalBind`. - -* In `Data.Vec.Base` the definitions `init` and `last` have been changed from the `initLast` - view-derived implementation to direct recursive definitions. - -* In `Data.Vec.Properties` the type of the proof `zipWith-comm` has been generalised from: - ```agda - zipWith-comm : ∀ {f : A → A → B} (comm : ∀ x y → f x y ≡ f y x) (xs ys : Vec A n) → - zipWith f xs ys ≡ zipWith f ys xs - ``` - to - ```agda - zipWith-comm : ∀ {f : A → B → C} {g : B → A → C} (comm : ∀ x y → f x y ≡ g y x) (xs : Vec A n) ys → - zipWith f xs ys ≡ zipWith g ys xs - ``` - -* In `Data.Vec.Relation.Unary.All` the functions `lookup` and `tabulate` have - been moved to `Data.Vec.Relation.Unary.All.Properties` and renamed - `lookup⁺` and `lookup⁻` respectively. - -* In `Data.Vec.Base` and `Data.Vec.Functional` the functions `iterate` and `replicate` - now take the length of vector, `n`, as an explicit rather than an implicit argument, i.e. - the new types are: - ```agda - iterate : (A → A) → A → ∀ n → Vec A n - replicate : ∀ n → A → Vec A n - ``` - -* In `Relation.Binary.Construct.Closure.Symmetric` the operation - `SymClosure` on relations in 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 `Relation.Nullary.Decidable.Core` the name `excluded-middle` has been - renamed to `¬¬-excluded-middle`. - - -Other 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` - -* Raw bundles by design don't contain any proofs so should in theory be able to live - in `Data.X.Base` rather than `Data.X.Bundles`. - -* To achieve this while keeping the dependency footprint small, the definitions of - raw bundles (`RawMagma`, `RawMonoid` etc.) have been moved from `Algebra(.Lattice)Bundles` to - a new module `Algebra(.Lattice).Bundles.Raw` which can be imported at much lower cost - from `Data.X.Base`. - -* We have then moved raw bundles defined in `Data.X.Properties` to `Data.X.Base` for - `X` = `Nat`/`Nat.Binary`/`Integer`/`Rational`/`Rational.Unnormalised`. - -Deprecated modules ------------------- - -### Moving `Data.Erased` to `Data.Irrelevant` - -* This fixes the fact we had picked the wrong name originally. The erased modality - corresponds to `@0` whereas the irrelevance one corresponds to `.`. - -### 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.Nat.Properties.Core` - -* The module `Data.Nat.Properties.Core` has been deprecated, and its one lemma moved to `Data.Nat.Base`, renamed as `s≤s⁻¹` - -### 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`. - -### Moving `Function.Related` - -* 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 - ``` - -### Moving `Relation.Binary.Construct.(Converse/Flip)` - -* The following files have been moved: - ```agda - Relation.Binary.Construct.Converse ↦ Relation.Binary.Construct.Flip.EqAndOrd - Relation.Binary.Construct.Flip ↦ Relation.Binary.Construct.Flip.Ord - ``` - -### Moving `Relation.Binary.Properties.XLattice` - -* The following files have been moved: - ```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.HeytingAlgebra.agda ↦ Relation.Binary.Lattice.Properties.HeytingAlgebra.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 - ``` - -Deprecated names ----------------- - -* In `Algebra.Construct.Zero`: - ```agda - rawMagma ↦ Algebra.Construct.Terminal.rawMagma - magma ↦ Algebra.Construct.Terminal.magma - semigroup ↦ Algebra.Construct.Terminal.semigroup - band ↦ Algebra.Construct.Terminal.band - ``` - -* In `Codata.Guarded.Stream.Properties`: - ```agda - map-identity ↦ map-id - map-fusion ↦ map-∘ - drop-fusion ↦ drop-drop - ``` - -* In `Codata.Sized.Colist.Properties`: - ```agda - map-identity ↦ map-id - map-map-fusion ↦ map-∘ - drop-drop-fusion ↦ drop-drop - ``` - -* In `Codata.Sized.Covec.Properties`: - ```agda - map-identity ↦ map-id - map-map-fusion ↦ map-∘ - ``` - -* In `Codata.Sized.Delay.Properties`: - ```agda - map-identity ↦ map-id - map-map-fusion ↦ map-∘ - map-unfold-fusion ↦ map-unfold - ``` - -* In `Codata.Sized.M.Properties`: - ```agda - map-compose ↦ map-∘ - ``` - -* In `Codata.Sized.Stream.Properties`: - ```agda - map-identity ↦ map-id - map-map-fusion ↦ map-∘ - ``` - -* In `Data.Container.Related`: - ```agda - _∼[_]_ ↦ _≲[_]_ - ``` - -* In `Data.Bool.Properties` (Issue #2046): - ```agda - push-function-into-if ↦ if-float - ``` - -* In `Data.Fin.Base`: two new, hopefully more memorable, names `↑ˡ` `↑ʳ` - for the 'left', resp. 'right' injection of a Fin m into a 'larger' type, - `Fin (m + n)`, resp. `Fin (n + m)`, with argument order to reflect the - position of the `Fin m` argument. - ``` - inject+ ↦ flip _↑ˡ_ - raise ↦ _↑ʳ_ - ``` - -* In `Data.Fin.Properties`: - ``` - toℕ-raise ↦ toℕ-↑ʳ - toℕ-inject+ ↦ toℕ-↑ˡ - splitAt-inject+ ↦ splitAt-↑ˡ m i n - splitAt-raise ↦ splitAt-↑ʳ - Fin0↔⊥ ↦ 0↔⊥ - eq? ↦ inj⇒≟ - ``` - -* In `Data.Fin.Permutation.Components`: - ``` - reverse ↦ Data.Fin.Base.opposite - reverse-prop ↦ Data.Fin.Properties.opposite-prop - reverse-involutive ↦ Data.Fin.Properties.opposite-involutive - reverse-suc ↦ Data.Fin.Properties.opposite-suc - ``` - -* In `Data.Integer.DivMod` the operator names have been renamed to - be consistent with those in `Data.Nat.DivMod`: - ``` - _divℕ_ ↦ _/ℕ_ - _div_ ↦ _/_ - _modℕ_ ↦ _%ℕ_ - _mod_ ↦ _%_ - ``` - -* In `Data.Integer.Properties` references to variables in names have - been made consistent so that `m`, `n` always refer to naturals and - `i` and `j` always refer to integers: - ``` - ≤-steps ↦ i≤j⇒i≤k+j - ≤-step ↦ i≤j⇒i≤1+j - - ≤-steps-neg ↦ i≤j⇒i-k≤j - ≤-step-neg ↦ i≤j⇒pred[i]≤j - - n≮n ↦ i≮i - ∣n∣≡0⇒n≡0 ↦ ∣i∣≡0⇒i≡0 - ∣-n∣≡∣n∣ ↦ ∣-i∣≡∣i∣ - 0≤n⇒+∣n∣≡n ↦ 0≤i⇒+∣i∣≡i - +∣n∣≡n⇒0≤n ↦ +∣i∣≡i⇒0≤i - +∣n∣≡n⊎+∣n∣≡-n ↦ +∣i∣≡i⊎+∣i∣≡-i - ∣m+n∣≤∣m∣+∣n∣ ↦ ∣i+j∣≤∣i∣+∣j∣ - ∣m-n∣≤∣m∣+∣n∣ ↦ ∣i-j∣≤∣i∣+∣j∣ - signₙ◃∣n∣≡n ↦ signᵢ◃∣i∣≡i - ◃-≡ ↦ ◃-cong - ∣m-n∣≡∣n-m∣ ↦ ∣i-j∣≡∣j-i∣ - m≡n⇒m-n≡0 ↦ i≡j⇒i-j≡0 - m-n≡0⇒m≡n ↦ i-j≡0⇒i≡j - m≤n⇒m-n≤0 ↦ i≤j⇒i-j≤0 - m-n≤0⇒m≤n ↦ i-j≤0⇒i≤j - m≤n⇒0≤n-m ↦ i≤j⇒0≤j-i - 0≤n-m⇒m≤n ↦ 0≤i-j⇒j≤i - n≤1+n ↦ i≤suc[i] - n≢1+n ↦ i≢suc[i] - m≤pred[n]⇒m-commute ↦ map-<∣> - ``` - -* In `Data.Nat.Properties`: - ```agda - 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>=_ (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) - infix 4 _ℕ<_ _ℕ≤infinity _ℕ≤_ (Codata.Sized.Conat) - infix 6 _ℕ+_ _+ℕ_ (Codata.Sized.Conat) - infixr 8 _⇒_ _⊸_ (Data.Container.Core) - infixr -1 _<$>_ _<*>_ (Data.Container.FreeMonad) - infixl 1 _>>=_ (Data.Container.FreeMonad) - infix 5 _▷_ (Data.Container.Indexed) - infixr 4 _,_ (Data.Container.Relation.Binary.Pointwise) - infix 4 _≈_ (Data.Float.Base) - 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) - infixl 7 _⊓′_ (Data.Nat.Base) - infixl 6 _⊔′_ (Data.Nat.Base) - infixr 8 _^_ (Data.Nat.Base) - infix 4 _!≢0 _!*_!≢0 (Data.Nat.Properties) - infixl 6.5 _P′_ _P_ _C′_ _C_ (Data.Nat.Combinatorics.Base) - infix 8 _⁻¹ (Data.Parity.Base) - 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) - infix 4 _≃?_ (Data.Rational.Unnormalised.Properties) - infixr 4 _,_ (Data.Refinement) - 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) - infixr 4 _,_ (Data.Tree.AVL.Value) - infix 4 _≈ₖᵥ_ (Data.Tree.AVL.Map.Membership.Propositional) - infixr 5 _`∷_ (Data.Vec.Reflection) - infixr 5 _∷=_ (Data.Vec.Membership.Setoid) - infix -1 _$ⁿ_ (Data.Vec.N-ary) - infix 4 _≋_ (Data.Vec.Functional.Relation.Binary.Equality.Setoid) - infixl 1 _>>=-cong_ _≡->>=-cong_ (Effect.Monad.Partiality) - infixl 1 _?>=′_ (Effect.Monad.Predicate) - infixl 1 _>>=-cong_ _>>=-congP_ (Effect.Monad.Partiality.All) - infixr 5 _∷_ (Foreign.Haskell.List.NonEmpty) - infixr 4 _,_ (Foreign.Haskell.Pair) - 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 4 _≈_ (Function.Metric.Rational.Bundles) - infix 3 _←_ _↢_ (Function.Related) - infix 4 _<_ (Induction.WellFounded) - infixl 6 _ℕ+_ (Level.Literals) - infixr 4 _,_ (Reflection.AnnotatedAST) - 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 4 _,_ (Reflection.AST.Traversal) - infix 4 _∈FV_ (Reflection.AST.DeBruijn) - infixr 9 _;_ (Relation.Binary.Construct.Composition) - infixl 6 _+²_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) - infixr -1 _atₛ_ (Relation.Binary.Indexed.Heterogeneous.Construct.At) - infixr -1 _atₛ_ (Relation.Binary.Indexed.Homogeneous.Construct.At) - infix 4 _∈_ _∉_ (Relation.Unary.Indexed) - 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 6 _∪_ (Relation.Binary.Construct.Union) - infixl 6 _+ℤ_ (Relation.Binary.HeterogeneousEquality.Quotients.Examples) - infix 4 _≉_ _≈ᵢ_ _≤ᵢ_ (Relation.Binary.Indexed.Homogeneous.Bundles) - infixr 9 _⍮_ (Relation.Unary.PredicateTransformer) - infix 8 ∼_ (Relation.Unary.PredicateTransformer) - infix 2 _×?_ _⊙?_ (Relation.Unary.Properties) - infix 10 _~? (Relation.Unary.Properties) - infixr 1 _⊎?_ (Relation.Unary.Properties) - infixr 7 _∩?_ (Relation.Unary.Properties) - infixr 6 _∪?_ (Relation.Unary.Properties) - infixr 5 _∷ᴹ_ _∷⁻¹ᴹ_ (Text.Regex.Search) - infixl 6 _`⊜_ (Tactic.RingSolver) - infix 8 ⊝_ (Tactic.RingSolver.Core.Expression) - infix 4 _∈ᴿ?_ _∉ᴿ?_ _∈?ε _∈?[_] _∈?[^_] (Text.Regex.Properties) - infix 4 _∈?_ _∉?_ (Text.Regex.Derivative.Brzozowski) - infix 4 _∈_ _∉_ _∈?_ _∉?_ (Text.Regex.String.Unsafe) - ``` - -New modules ------------ - -* Constructive algebraic structures with apartness relations: - ``` - Algebra.Apartness - Algebra.Apartness.Bundles - Algebra.Apartness.Structures - Algebra.Apartness.Properties.CommutativeHeytingAlgebra - Relation.Binary.Properties.ApartnessRelation - ``` - -* Algebraic structures obtained by flipping their binary operations: - ``` - Algebra.Construct.Flip.Op - ``` - -* Algebraic structures when freely adding an identity element: - ``` - Algebra.Construct.Add.Identity - ``` - -* Definitions for algebraic modules: - ``` - Algebra.Module - Algebra.Module.Core - Algebra.Module.Definitions.Bi.Simultaneous - Algebra.Module.Morphism.Construct.Composition - Algebra.Module.Morphism.Construct.Identity - Algebra.Module.Morphism.Definitions - Algebra.Module.Morphism.ModuleHomomorphism - Algebra.Module.Morphism.Structures - Algebra.Module.Properties - ``` - -* Identity morphisms and composition of morphisms between algebraic structures: - ``` - Algebra.Morphism.Construct.Composition - Algebra.Morphism.Construct.Identity - ``` - -* Properties of various new algebraic structures: - ``` - Algebra.Properties.MoufangLoop - Algebra.Properties.Quasigroup - Algebra.Properties.MiddleBolLoop - Algebra.Properties.Loop - Algebra.Properties.KleeneAlgebra - ``` - -* Properties of rings without a unit - ``` - Algebra.Properties.RingWithoutOne` - ``` - -* Proof of the Binomial Theorem for semirings - ``` - Algebra.Properties.Semiring.Binomial - Algebra.Properties.CommutativeSemiring.Binomial - ``` - -* 'Optimised' tail-recursive exponentiation properties: - ``` - Algebra.Properties.Semiring.Exp.TailRecursiveOptimised - ``` - -* An implementation of M-types with `--guardedness` flag: - ``` - Codata.Guarded.M - ``` - -* An implementation of M-types with `--guardedness` flag: - ``` - Codata.Guarded.M - ``` - -* A definition of infinite streams using coinductive records: - ``` - Codata.Guarded.Stream - Codata.Guarded.Stream.Properties - Codata.Guarded.Stream.Relation.Binary.Pointwise - Codata.Guarded.Stream.Relation.Unary.All - Codata.Guarded.Stream.Relation.Unary.Any - ``` - -* A small library for function arguments with default values: - ``` - Data.Default - ``` - -* A small library defining a structurally recursive view of `Fin n`: - ``` - Data.Fin.Relation.Unary.Top - ``` - -* A small library for a non-empty fresh list: - ``` - Data.List.Fresh.NonEmpty - ``` - -* A small library defining a structurally inductive view of lists: - ``` - Data.List.Relation.Unary.Sufficient - ``` - -* Combinations and permutations for ℕ. - ``` - Data.Nat.Combinatorics - Data.Nat.Combinatorics.Base - Data.Nat.Combinatorics.Spec - ``` - -* A small library defining parity and its algebra: - ``` - Data.Parity - Data.Parity.Base - Data.Parity.Instances - Data.Parity.Properties - ``` - -* New base module for `Data.Product` containing only the basic definitions. - ``` - Data.Product.Base - ``` - -* Reflection utilities for some specific types: - ``` - Data.List.Reflection - Data.Vec.Reflection - ``` - -* The `All` predicate over non-empty lists: - ``` - Data.List.NonEmpty.Relation.Unary.All - ``` - -* Some n-ary functions manipulating lists - ``` - Data.List.Nary.NonDependent - ``` - -* Added Logarithm base 2 on natural numbers: - ``` - Data.Nat.Logarithm.Core - Data.Nat.Logarithm - ``` - -* Show module for unnormalised rationals: - ``` - Data.Rational.Unnormalised.Show - ``` - -* Membership relations for maps and sets - ``` - Data.Tree.AVL.Map.Membership.Propositional - Data.Tree.AVL.Map.Membership.Propositional.Properties - Data.Tree.AVL.Sets.Membership - Data.Tree.AVL.Sets.Membership.Properties - ``` - -* Port of `Linked` to `Vec`: - ``` - Data.Vec.Relation.Unary.Linked - Data.Vec.Relation.Unary.Linked.Properties - ``` - -* Combinators for propositional equational reasoning on vectors with different indices - ``` - Data.Vec.Relation.Binary.Equality.Cast - ``` - -* Relations on indexed sets - ``` - Function.Indexed.Bundles - ``` - -* Properties of various types of functions: - ``` - Function.Consequences - Function.Properties.Bijection - Function.Properties.RightInverse - Function.Properties.Surjection - Function.Construct.Constant - ``` - -* New interface for `NonEmpty` Haskell lists: - ``` - Foreign.Haskell.List.NonEmpty - ``` - -* The 'no infinite descent' principle for (accessible elements of) well-founded relations: - ``` - Induction.NoInfiniteDescent - ``` - -* In order to improve modularity, the contents of `Relation.Binary.Lattice` has been - split out into the standard: - ``` - Relation.Binary.Lattice.Definitions - Relation.Binary.Lattice.Structures - Relation.Binary.Lattice.Bundles - ``` - All contents is re-exported by `Relation.Binary.Lattice` as before. - - -* Added relational reasoning over apartness relations: - ``` - Relation.Binary.Reasoning.Base.Apartness` - ``` - -* Algebraic properties of `_∩_` and `_∪_` for predicates - ``` - Relation.Unary.Algebra - ``` - -* Both versions of equality on predicates are equivalences - ``` - Relation.Unary.Relation.Binary.Equality - ``` - -* The subset relations on predicates define an order - ``` - Relation.Unary.Relation.Binary.Subset - ``` - -* Polymorphic versions of some unary relations and their properties - ``` - Relation.Unary.Polymorphic - Relation.Unary.Polymorphic.Properties - ``` - -* Alpha equality over reflected terms - ``` - Reflection.AST.AlphaEquality - ``` - -* Various system types and primitives: - ``` - System.Clock.Primitive - System.Clock - System.Console.ANSI - System.Directory.Primitive - System.Directory - System.FilePath.Posix.Primitive - System.FilePath.Posix - System.Process.Primitive - System.Process - ``` - -* A new `cong!` tactic for automatically deriving arguments to `cong` - ``` - Tactic.Cong - ``` - -* A golden testing library with test pools, an options parser, a runner: - ``` - Test.Golden - ``` - -Additions to existing modules ------------------------------ - -* 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 Quasigroup 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 Quasiring c ℓ : Set (suc (c ⊔ ℓ)) - record Nearring c ℓ : Set (suc (c ⊔ ℓ)) - 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 ⊔ ℓ)) - ``` - -* Added new definitions to `Algebra.Bundles.Raw`: - ```agda - record RawLoop c ℓ : Set (suc (c ⊔ ℓ)) - record RawQuasiGroup c ℓ : Set (suc (c ⊔ ℓ)) - record RawRingWithoutOne c ℓ : Set (suc (c ⊔ ℓ)) - ``` - -* 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 ⊔ ℓ) - ``` - -* Added new proof 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 _+_ = ∀ w x y z → ((w + x) * (y + z)) ≈ ((w + y) * (x + z)) - - SelfInverse f = ∀ {x y} → f x ≈ y → f y ≈ x - - LeftDividesˡ _∙_ _\\_ = ∀ x y → (x ∙ (x \\ y)) ≈ y - LeftDividesʳ _∙_ _\\_ = ∀ x y → (x \\ (x ∙ y)) ≈ y - RightDividesˡ _∙_ _//_ = ∀ x y → ((y // x) ∙ x) ≈ y - RightDividesʳ _∙_ _//_ = ∀ x y → ((y ∙ x) // x) ≈ y - LeftDivides ∙ \\ = (LeftDividesˡ ∙ \\) × (LeftDividesʳ ∙ \\) - RightDivides ∙ // = (RightDividesˡ ∙ //) × (RightDividesʳ ∙ //) - - 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 - ``` - -* In `Algebra.Bundles.Lattice` the existing record `Lattice` now provides - ```agda - ∨-commutativeSemigroup : CommutativeSemigroup c ℓ - ∧-commutativeSemigroup : CommutativeSemigroup c ℓ - ``` - and their corresponding algebraic sub-bundles. - -* In `Algebra.Lattice.Structures` the record `IsLattice` now provides - ``` - ∨-isCommutativeSemigroup : IsCommutativeSemigroup ∨ - ∧-isCommutativeSemigroup : IsCommutativeSemigroup ∧ - ``` - and their corresponding algebraic substructures. - -* The module `Algebra.Properties.Magma.Divisibility` now re-exports operations - `_∣ˡ_`, `_∤ˡ_`, `_∣ʳ_`, `_∤ʳ_` from `Algebra.Definitions.Magma`. - -* 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 to `Algebra.Properties.CommutativeSemigroup`: - ```agda - xy∙xx≈x∙yxx : (x ∙ y) ∙ (x ∙ x) ≈ x ∙ (y ∙ (x ∙ x)) - - interchange : Interchangable _∙_ _∙_ - leftSemimedial : LeftSemimedial _∙_ - rightSemimedial : RightSemimedial _∙_ - middleSemimedial : (x ∙ y) ∙ (z ∙ x) ≈ (x ∙ z) ∙ (y ∙ x) - semimedial : Semimedial _∙_ - ``` - -* 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 proof to `Algebra.Properties.Monoid.Mult`: - ```agda - ×-congˡ : (_× x) Preserves _≡_ ⟶ _≈_ - ``` - -* Added new proof to `Algebra.Properties.Monoid.Sum`: - ```agda - sum-init-last : (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 : - 1# * x ≈ - x - x+x≈x⇒x≈0 : x + x ≈ x → x ≈ 0# - x[y-z]≈xy-xz : x * (y - z) ≈ x * y - x * z - [y-z]x≈yx-zx : (y - z) * x ≈ (y * x) - (z * x) - ``` - -* 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 → concat ass ≈ concat bss - map-concat : map f (concat ass) ≈ concat (map (List⁺.map f) ass) - lookup-transpose : lookup n (transpose ass) ≡ List.map (lookup n) ass - lookup-transpose⁺ : lookup n (transpose⁺ ass) ≡ List⁺.map (lookup n) ass - ``` - -* 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 - ``` - -* Exposed container combinator conversion functions from `Data.Container.Combinator.Properties` - separately 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 → ⟦ const A ⟧ B - from-const : ⟦ const A ⟧ B → A - to-∘ : ⟦ C₁ ⟧ (⟦ C₂ ⟧ A) → ⟦ C₁ ∘ C₂ ⟧ A - from-∘ : ⟦ C₁ ∘ C₂ ⟧ A → ⟦ C₁ ⟧ (⟦ C₂ ⟧ A) - to-× : ⟦ C₁ ⟧ A P.× ⟦ C₂ ⟧ A → ⟦ C₁ × C₂ ⟧ A - from-× : ⟦ C₁ × C₂ ⟧ A → ⟦ C₁ ⟧ A P.× ⟦ C₂ ⟧ A - to-Π : (∀ i → ⟦ Cᵢ i ⟧ A) → ⟦ Π I Cᵢ ⟧ A - from-Π : ⟦ Π I Cᵢ ⟧ A → ∀ i → ⟦ Cᵢ i ⟧ A - to-⊎ : ⟦ C₁ ⟧ A S.⊎ ⟦ C₂ ⟧ A → ⟦ C₁ ⊎ C₂ ⟧ A - from-⊎ : ⟦ C₁ ⊎ C₂ ⟧ A → ⟦ C₁ ⟧ A S.⊎ ⟦ C₂ ⟧ A - to-Σ : (∃ λ i → ⟦ C i ⟧ A) → ⟦ Σ I C ⟧ A - from-Σ : ⟦ Σ I C ⟧ A → ∃ λ i → ⟦ C i ⟧ A - ``` - -* 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`: - ```agda - spo-wellFounded : IsStrictPartialOrder _≈_ _⊏_ → WellFounded _⊏_ - spo-noetherian : IsStrictPartialOrder _≈_ _⊏_ → WellFounded (flip _⊏_) - - <-weakInduction-startingFrom : P i → (∀ j → P (inject₁ j) → P (suc j)) → ∀ {j} → j ≥ i → P j - ``` - -* 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 π) ≈ π - ``` - -* Added new proofs in `Data.Fin.Properties`: - ``` - 1↔⊤ : Fin 1 ↔ ⊤ - 2↔Bool : Fin 2 ↔ Bool - - 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′ - i≤inject₁[j]⇒i≤1+j : i ≤ inject₁ j → i ≤ suc j - ``` - -* Added new lemmas in `Data.Fin.Substitution.Lemmas.TermLemmas`: - ``` - map-var≡ : (∀ x → lookup ρ₁ x ≡ f x) → (∀ x → lookup ρ₂ x ≡ T.var (f x)) → map var ρ₁ ≡ ρ₂ - wk≡wk : map var VarSubst.wk ≡ T.wk {n = n} - id≡id : map var VarSubst.id ≡ T.id {n = n} - sub≡sub : map var (VarSubst.sub x) ≡ T.sub (T.var x) - ↑≡↑ : map var (ρ VarSubst.↑) ≡ map T.var ρ T.↑ - /Var≡/ : t /Var ρ ≡ t T./ map T.var ρ - - sub-renaming-commutes : t /Var VarSubst.sub x T./ ρ ≡ t T./ ρ T.↑ T./ T.sub (lookup ρ x) - sub-commutes-with-renaming : t T./ T.sub t′ /Var ρ ≡ t /Var ρ VarSubst.↑ T./ T.sub (t′ /Var ρ) - ``` - -* Added new functions and definitions in `Data.Integer.Base`: - ```agda - _^_ : ℤ → ℕ → ℤ - - +-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 and definitions to `Data.List.Base`: - ```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) - - catMaybes : List (Maybe A) → List A - ap : List (A → B) → List A → List B - ++-rawMagma : Set a → RawMagma a _ - ++-[]-rawMonoid : Set a → RawMonoid a _ - - iterate : (A → A) → A → ℕ → List A - insertAt : (xs : List A) → Fin (suc (length xs)) → A → List A - updateAt : (xs : List A) → Fin (length xs) → (A → A) → List 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 - ``` - -* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`: - ```agda - foldr-commMonoid : xs ↭ ys → foldr _∙_ ε xs ≈ foldr _∙_ ε ys - ``` - -* Added new function to `Data.List.Relation.Binary.Permutation.Propositional.Properties` - ```agda - ↭-reverse : reverse xs ↭ xs - ``` - -* Added new proofs to `Data.List.Relation.Binary.Sublist.Setoid.Properties`: - ``` - ⊆-mergeˡ : xs ⊆ merge _≤?_ xs ys - ⊆-mergeʳ : ys ⊆ merge _≤?_ xs ys - ``` - -* Added new functions in `Data.List.Relation.Unary.All`: - ``` - decide : Π[ P ∪ Q ] → Π[ All P ∪ Any Q ] - ``` - -* Added new proof to `Data.List.Relation.Unary.All.Properties`: - ```agda - gmap⁻ : Q ∘ f ⋐ P → All Q ∘ map f ⋐ All P - ``` - -* 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₂ - - ∈-++⁺∘++⁻ : (p : v ∈ xs ++ ys) → [ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ (∈-++⁻ xs p) ≡ p - ∈-++⁻∘++⁺ : (p : v ∈ xs ⊎ v ∈ ys) → ∈-++⁻ xs ([ ∈-++⁺ˡ , ∈-++⁺ʳ xs ]′ p) ≡ p - ∈-++↔ : (v ∈ xs ⊎ v ∈ ys) ↔ v ∈ xs ++ ys - ∈-++-comm : v ∈ xs ++ ys → v ∈ ys ++ xs - ∈-++-comm∘++-comm : (p : v ∈ xs ++ ys) → ∈-++-comm ys xs (∈-++-comm xs ys p) ≡ p - ∈-++↔++ : v ∈ xs ++ ys ↔ v ∈ ys ++ xs - ``` - -* 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 - drop-all : n ≥ length xs → drop n xs ≡ [] - - take-[] : take m [] ≡ [] - drop-[] : drop m [] ≡ [] - - drop-drop : drop n (drop m xs) ≡ drop (m + n) xs - - lookup-replicate : lookup (replicate n x) i ≡ x - map-replicate : map f (replicate n x) ≡ replicate n (f x) - zipWith-replicate : zipWith _⊕_ (replicate n x) (replicate n y) ≡ replicate n (x ⊕ y) - - length-iterate : length (iterate f x n) ≡ n - iterate-id : iterate id x n ≡ replicate n x - lookup-iterate : lookup (iterate f x n) (cast (sym (length-iterate f x n)) i) ≡ ℕ.iterate f x (toℕ i) - - length-insertAt : length (insertAt xs i v) ≡ suc (length xs) - length-removeAt′ : length xs ≡ suc (length (removeAt xs k)) - removeAt-insertAt : removeAt (insertAt xs i v) ((cast (sym (length-insertAt xs i v)) i)) ≡ xs - insertAt-removeAt : insertAt (removeAt xs i) (cast (sym (lengthAt-removeAt xs i)) i) (lookup xs i) ≡ xs - - 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 - ``` - -* In `Data.List.NonEmpty.Base`: - ```agda - drop+ : ℕ → List⁺ A → List⁺ A - ``` -* 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 proof to `Data.Maybe.Properties` - ```agda - <∣>-idem : Idempotent _<∣>_ - ``` - -* 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 - - s0 : .{{_ : NonZero m}} n → m ^ n > 0 - - ^-monoˡ-≤ : (_^ n) Preserves _≤_ ⟶ _≤_ - ^-monoʳ-≤ : .{{_ : NonZero m}} → (m ^_) Preserves _≤_ ⟶ _≤_ - ^-monoˡ-< : .{{_ : NonZero n}} → (_^ n) Preserves _<_ ⟶ _<_ - ^-monoʳ-< : 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) - - <-dense : Dense _<_ - <-isDenseLinearOrder : IsDenseLinearOrder _≡_ _<_ - <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ - ``` - -* Added new rounding functions in `Data.Rational.Unnormalised.Base`: - ```agda - floor ceiling truncate round ⌊_⌋ ⌈_⌉ [_] : ℚᵘ → ℤ - fracPart : ℚᵘ → ℚᵘ - ``` - -* Added new definitions in `Data.Rational.Unnormalised.Properties`: - ```agda - ↥p≡0⇒p≃0 : ↥ p ≡ 0ℤ → p ≃ 0ℚᵘ - p≃0⇒↥p≡0 : p ≃ 0ℚᵘ → ↥ p ≡ 0ℤ - ↥p≡↥q≡0⇒p≃q : ↥ p ≡ 0ℤ → ↥ q ≡ 0ℤ → p ≃ q - - +-*-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) - - 0≄1 : 0ℚᵘ ≄ 1ℚᵘ - ≃-≄-irreflexive : Irreflexive _≃_ _≄_ - ≄-symmetric : Symmetric _≄_ - ≄-cotransitive : Cotransitive _≄_ - ≄⇒invertible : p ≄ q → Invertible _≃_ 1ℚᵘ _*_ (p - q) - - <-dense : Dense _<_ - - <-isDenseLinearOrder : IsDenseLinearOrder _≃_ _<_ - +-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ - +-*-isHeytingField : IsHeytingField _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ - - <-denseLinearOrder : DenseLinearOrder 0ℓ 0ℓ 0ℓ - +-*-heytingCommutativeRing : HeytingCommutativeRing 0ℓ 0ℓ 0ℓ - +-*-heytingField : HeytingField 0ℓ 0ℓ 0ℓ - - module ≃-Reasoning = SetoidReasoning ≃-setoid - ``` - -* 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`: - ``` - Σ-≡,≡→≡ : (∃ λ (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`: - ```agda - ≤-isDecTotalOrder-≈ : IsDecTotalOrder _≈_ _≤_ - ≤-decTotalOrder-≈ : DecTotalOrder _ _ _ - ``` -* Added new definitions in `Data.Sum.Properties`: - ```agda - swap-↔ : (A ⊎ B) ↔ (B ⊎ A) - ``` - -* Added new proofs in `Data.Sum.Properties`: - ```agda - map-assocˡ : map (map f g) h ∘ assocˡ ≗ assocˡ ∘ map f (map g h) - map-assocʳ : 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 = A → B n → B (suc n) - FoldlOp A B = 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 : (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) - ∷ʳ-++ : cast eq ((xs ∷ʳ a) ++ ys) ≡ xs ++ (a ∷ ys) - - 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 - cast-++ˡ : cast (cong (_+ n) eq) (xs ++ ys) ≡ cast eq xs ++ ys - cast-++ʳ : cast (cong (m +_) eq) (xs ++ ys) ≡ xs ++ cast eq ys - - iterate-id : iterate id x n ≡ replicate x - take-iterate : take n (iterate f x (n + m)) ≡ iterate f x n - drop-iterate : drop n (iterate f x n) ≡ [] - lookup-iterate : lookup (iterate f x n) i ≡ ℕ.iterate f x (toℕ i) - toList-iterate : toList (iterate f x n) ≡ List.iterate f x n - - 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 - fromList-reverse : cast (Listₚ.length-reverse xs) (fromList (List.reverse xs)) ≡ reverse (fromList xs) - - ∷-ʳ++ : cast eq ((a ∷ xs) ʳ++ ys) ≡ xs ʳ++ (a ∷ ys) - ++-ʳ++ : cast eq ((xs ++ ys) ʳ++ zs) ≡ ys ʳ++ (xs ʳ++ zs) - ʳ++-ʳ++ : cast eq ((xs ʳ++ ys) ʳ++ zs) ≡ ys ʳ++ (xs ++ zs) - - length-toList : List.length (toList xs) ≡ length xs - toList-insertAt : toList (insertAt xs i v) ≡ List.insertAt (toList xs) (Fin.cast (cong suc (sym (length-toList xs))) i) v - - 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 isomorphisms to `Data.Unit.Polymorphic.Properties`: - ```agda - ⊤↔⊤* : ⊤ ↔ ⊤* - ``` - -* 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 < [] - - <-respectsˡ : IsPartialEquivalence _≈_ → _≺_ Respectsˡ _≈_ → _<_ Respectsˡ _≋_ - <-respectsʳ : IsPartialEquivalence _≈_ → _≺_ Respectsʳ _≈_ → _<_ _Respectsʳ _≋_ - <-wellFounded : Transitive _≈_ → _≺_ Respectsʳ _≈_ → WellFounded _≺_ → WellFounded _<_ - ``` - -* 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 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 ] - - 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) - ``` - -* Added vector associativity proof to `Data.Vec.Relation.Binary.Equality.Setoid`: - ``` - ++-assoc : (xs ++ ys) ++ zs ≋ xs ++ (ys ++ zs) - ``` - -* 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↔ : A ↔ B → A ^ n ↔ B ^ n - Fin[m^n]↔Fin[m]^n : Fin (m ^ n) ↔ Fin m Vec.^ n - ``` - -* 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 a non-dependent version of `flip` in `Function.Base`: - ```agda - flip′ : (A → B → C) → (B → A → C) - ``` - -* Added new proofs and definitions in `Function.Bundles`: - ```agda - LeftInverse.isSplitSurjection : LeftInverse → IsSplitSurjection to - LeftInverse.surjection : LeftInverse → Surjection - SplitSurjection = LeftInverse - _⟨$⟩_ = Func.to - ``` - -* Added new proofs to `Function.Properties.Inverse`: - ```agda - ↔-refl : Reflexive _↔_ - ↔-sym : Symmetric _↔_ - ↔-trans : Transitive _↔_ - - ↔⇒↣ : A ↔ B → A ↣ B - ↔-fun : A ↔ B → C ↔ D → (A → C) ↔ (B → D) - - Inverse⇒Injection : Inverse S T → Injection S T - ``` - -* 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 proof and record in `Function.Structures`: - ```agda - record IsSplitSurjection (f : A → B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) - ``` - -* Added new definition to the `Surjection` module in `Function.Related.Surjection`: - ``` - f⁻ = proj₁ ∘ surjective - ``` - -* Added new proof to `Induction.WellFounded` - ```agda - Acc-resp-flip-≈ : _<_ Respectsʳ (flip _≈_) → (Acc _<_) Respects _≈_ - acc⇒asym : Acc _<_ x → x < y → ¬ (y < x) - wf⇒asym : WellFounded _<_ → Asymmetric _<_ - wf⇒irrefl : _<_ Respects₂ _≈_ → Symmetric _≈_ → WellFounded _<_ → Irreflexive _≈_ _<_ - ``` - -* 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 types and operations to `Reflection.TCM`: - ``` - Blocker : Set - blockerMeta : Meta → Blocker - blockerAny : List Blocker → Blocker - blockerAll : List Blocker → Blocker - blockTC : Blocker → TC A - ``` - -* 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 proof in `Relation.Binary.Construct.Closure.Transitive`: - ``` - transitive⁻ : Transitive _∼_ → TransClosure _∼_ ⇒ _∼_ - ``` - -* 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 _ - Stable : Pred A ℓ → Set _ - ``` - -* 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 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`: - ``` - RightTrans R S = Trans R S R - LeftTrans S R = Trans S R R - - Dense _<_ = x < y → ∃[ z ] x < z × z < y - Cotransitive _#_ = x # y → ∀ z → (x # z) ⊎ (z # y) - Tight _≈_ _#_ = (¬ 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 = (f x ⊑ y → x ≤ g y) × (x ≤ g y → f x ⊑ y) - ``` - -* Added new definitions in `Relation.Binary.Bundles`: - ``` - record DenseLinearOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - record ApartnessRelation c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where - ``` - -* Added new definitions in `Relation.Binary.Structures`: - ``` - record IsDenseLinearOrder (_<_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where - 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⁻ : Acc _∼⁺_ x → Acc _∼_ x - wellFounded⁻ : WellFounded _∼⁺_ → WellFounded _∼_ - accessible : 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 proof in `Relation.Nullary.Reflects`: - ```agda - T-reflects : Reflects (T b) b - T-reflects-elim : Reflects (T a) b → b ≡ a - ``` - -* Added new operations in `System.Exit`: - ``` - isSuccess : ExitCode → Bool - isFailure : ExitCode → Bool - ``` - -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 - ``` From 0c2288f5b5be6c140da3269995e096365809d9fc Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 17 Dec 2023 11:20:01 +0000 Subject: [PATCH 49/57] resolved merge conflict --- CHANGELOG.md | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 61ac7423af..51c4f49374 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -37,6 +37,7 @@ New modules Additions to existing modules ----------------------------- +<<<<<<< HEAD * In `Data.Fin.Properties`: ```agda nonZeroIndex : Fin n → ℕ.NonZero n @@ -85,4 +86,13 @@ Additions to existing modules pred-injective : .{{NonZero m}} → .{{NonZero n}} → pred m ≡ pred n → m ≡ n pred-cancel-≡ : pred m ≡ pred n → ((m ≡ 0 × n ≡ 1) ⊎ (m ≡ 1 × n ≡ 0)) ⊎ m ≡ n ``` - \ No newline at end of file + +* Added new definition in `Relation.Binary.Construct.Closure.Transitive` + ``` + transitive⁻ : Transitive _∼_ → TransClosure _∼_ ⇒ _∼_ + ``` + +* Added new definition in `Relation.Unary` + ``` + Stable : Pred A ℓ → Set _ + ``` From d074a0df87bd3e4fc044070dda00bf850a0cfde7 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 17 Dec 2023 11:21:54 +0000 Subject: [PATCH 50/57] resolved merge conflict, this time --- CHANGELOG.md | 1 - src/Relation/Binary/Construct/Closure/Transitive.agda | 7 ++----- 2 files changed, 2 insertions(+), 6 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 51c4f49374..582ca725ff 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -37,7 +37,6 @@ New modules Additions to existing modules ----------------------------- -<<<<<<< HEAD * In `Data.Fin.Properties`: ```agda nonZeroIndex : Fin n → ℕ.NonZero n diff --git a/src/Relation/Binary/Construct/Closure/Transitive.agda b/src/Relation/Binary/Construct/Closure/Transitive.agda index 9741d921a0..140d1d4a9d 100644 --- a/src/Relation/Binary/Construct/Closure/Transitive.agda +++ b/src/Relation/Binary/Construct/Closure/Transitive.agda @@ -71,11 +71,8 @@ module _ (_∼_ : Rel A ℓ) where transitive = _++_ transitive⁻ : Transitive _∼_ → _∼⁺_ ⇒ _∼_ - transitive⁻ trans = ∼⁺⊆∼ - where - ∼⁺⊆∼ : _∼⁺_ ⇒ _∼_ - ∼⁺⊆∼ [ x∼y ] = x∼y - ∼⁺⊆∼ (x∼y ∷ x∼⁺y) = trans x∼y (∼⁺⊆∼ x∼⁺y) + transitive⁻ trans [ x∼y ] = x∼y + transitive⁻ trans (x∼y ∷ x∼⁺y) = trans x∼y (transitive⁻ trans x∼⁺y) accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x accessible⁻ = ∼⊆∼⁺.accessible From 5e66965fed34c318769320385ca6d69cdea1c8d6 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Sun, 31 Dec 2023 17:36:45 +0000 Subject: [PATCH 51/57] revised in line with review comments --- src/Induction/NoInfiniteDescent.agda | 118 +++++++++++++++------------ 1 file changed, 66 insertions(+), 52 deletions(-) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 54f2b89710..4ccdb1fec4 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -32,31 +32,57 @@ private ------------------------------------------------------------------------ -- Definitions +module _ (f : ℕ → A) where + + module _ (_<_ : Rel A r) where + + InfiniteDescendingSequence : Set _ + InfiniteDescendingSequence = ∀ n → f (suc n) < f n + + InfiniteSequenceFrom : Pred A _ + InfiniteSequenceFrom x = f zero ≡ x × InfiniteDescendingSequence + + module _ (_<_ : Rel A r) where + + private + + _<⁺_ = TransClosure _<_ + + InfiniteDescendingSequence⁺ : Set _ + InfiniteDescendingSequence⁺ = ∀ {m n} → m ℕ.< n → f n <⁺ f m + + InfiniteSequence⁺From : Pred A _ + InfiniteSequence⁺From x = f zero ≡ x × InfiniteDescendingSequence⁺ + + sequence⁺ : InfiniteDescendingSequence _<⁺_ → InfiniteDescendingSequence⁺ + sequence⁺ seq[f] = seq⁺[f]′ ∘ ℕ.<⇒<′ + where + seq⁺[f]′ : ∀ {m n} → m ℕ.<′ n → f n <⁺ f m + seq⁺[f]′ ℕ.<′-base = seq[f] _ + seq⁺[f]′ (ℕ.<′-step m<′n) = seq[f] _ ++ seq⁺[f]′ m<′n + + sequence⁻ : InfiniteDescendingSequence⁺ → InfiniteDescendingSequence _<⁺_ + sequence⁻ seq[f] = seq[f] ∘ n<1+n + module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where DescentAt : Pred A _ DescentAt x = P x → ∃[ y ] y < x × P y Acc⇒Descent : Set _ - Acc⇒Descent = ∀[ Acc _<_ ⇒ DescentAt ] + Acc⇒Descent = ∀ {x} → Acc _<_ x → DescentAt x Descent : Set _ - Descent = ∀[ DescentAt ] - - InfiniteDescendingSequence : (f : ℕ → A) → Set _ - InfiniteDescendingSequence f = ∀ n → f (suc n) < f n - - InfiniteSequence_From_ : (f : ℕ → A) → Pred A _ - InfiniteSequence f From x = f zero ≡ x × InfiniteDescendingSequence f + Descent = ∀ {x} → DescentAt x InfiniteDescentAt : Pred A _ - InfiniteDescentAt x = P x → ∃[ f ] Π[ P ∘ f ] × InfiniteSequence f From x + InfiniteDescentAt x = P x → ∃[ f ] InfiniteSequenceFrom f _<_ x × ∀ z → P (f z) Acc⇒InfiniteDescent : Set _ - Acc⇒InfiniteDescent = ∀[ Acc _<_ ⇒ InfiniteDescentAt ] + Acc⇒InfiniteDescent = ∀ {x} → Acc _<_ x → InfiniteDescentAt x InfiniteDescent : Set _ - InfiniteDescent = ∀[ InfiniteDescentAt ] + InfiniteDescent = ∀ {x} → InfiniteDescentAt x ------------------------------------------------------------------------ -- Basic lemmas: assume unrestricted descent @@ -69,39 +95,42 @@ module InfiniteDescent (_<_ : Rel A r) (P : Pred A ℓ) where rec : _ rec y rec[y] py with z , z Date: Sun, 31 Dec 2023 18:58:28 +0000 Subject: [PATCH 52/57] tweaked `export`s; does that fix things now? --- src/Induction/NoInfiniteDescent.agda | 13 ++++++++----- 1 file changed, 8 insertions(+), 5 deletions(-) diff --git a/src/Induction/NoInfiniteDescent.agda b/src/Induction/NoInfiniteDescent.agda index 4ccdb1fec4..928b26f0a6 100644 --- a/src/Induction/NoInfiniteDescent.agda +++ b/src/Induction/NoInfiniteDescent.agda @@ -22,7 +22,7 @@ open import Relation.Binary.Construct.Closure.Transitive open import Relation.Binary.PropositionalEquality.Core open import Relation.Nullary.Negation.Core as Negation using (¬_) open import Relation.Unary - using (Pred; ∁; _∩_; _⇒_; Universal; IUniversal; Stable) + using (Pred; _∩_; _⇒_; Universal; IUniversal; Stable) private variable @@ -216,15 +216,18 @@ module FurtherCorollaries (_<_ : Rel A r) (P : Pred A ℓ) where module _ (_<_ : Rel A r) {P : Pred A ℓ} (stable : Stable P) where - open FurtherCorollaries _<_ (∁ P) - using (acc⇒noInfiniteDescent⁺; wf⇒noInfiniteDescent⁺) + private + module FC = FurtherCorollaries _<_ (¬_ ∘ P) + + open FC public + using () renaming (Acc⇒Descent⁺ to NoSmallestCounterExample) acc⇒noSmallestCounterExample : NoSmallestCounterExample → ∀ {x} → Acc _<_ x → P x - acc⇒noSmallestCounterExample noSmallest = stable _ ∘ acc⇒noInfiniteDescent⁺ noSmallest + acc⇒noSmallestCounterExample noSmallest = stable _ ∘ FC.acc⇒noInfiniteDescent⁺ noSmallest wf⇒noSmallestCounterExample : WellFounded _<_ → NoSmallestCounterExample → ∀ {x} → P x - wf⇒noSmallestCounterExample wf noSmallest = stable _ (wf⇒noInfiniteDescent⁺ noSmallest wf) + wf⇒noSmallestCounterExample wf noSmallest = stable _ (FC.wf⇒noInfiniteDescent⁺ noSmallest wf) ------------------------------------------------------------------------ -- Exports From 45188583ebc0c545167a522a3883066f5e03cf17 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Sun, 25 Feb 2024 19:37:14 +0800 Subject: [PATCH 53/57] Fix merge mistake --- src/Relation/Unary.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Relation/Unary.agda b/src/Relation/Unary.agda index 7b7807e486..9f3ea16633 100644 --- a/src/Relation/Unary.agda +++ b/src/Relation/Unary.agda @@ -15,7 +15,7 @@ open import Data.Sum.Base using (_⊎_; [_,_]) open import Function.Base using (_∘_; _|>_) open import Level using (Level; _⊔_; 0ℓ; suc; Lift) open import Relation.Nullary.Decidable.Core using (Dec; True) -open import Relation.Nullary.Negation.Core as Nullary using (¬_) +open import Relation.Nullary as Nullary using (¬_) open import Relation.Binary.PropositionalEquality.Core using (_≡_) private From f84d319e707c9e346f7fb4b5b1dc62da8ae80751 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Sun, 25 Feb 2024 20:59:06 +0800 Subject: [PATCH 54/57] Refactored to remove a indirection and nested modules --- src/Induction/InfiniteDescent.agda | 187 +++++++++++++++++++++ src/Induction/NoInfiniteDescent.agda | 238 --------------------------- 2 files changed, 187 insertions(+), 238 deletions(-) create mode 100644 src/Induction/InfiniteDescent.agda delete mode 100644 src/Induction/NoInfiniteDescent.agda diff --git a/src/Induction/InfiniteDescent.agda b/src/Induction/InfiniteDescent.agda new file mode 100644 index 0000000000..e0bd372102 --- /dev/null +++ b/src/Induction/InfiniteDescent.agda @@ -0,0 +1,187 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- A standard consequence of accessibility/well-foundedness: +-- the impossibility of 'infinite descent' from any (accessible) +-- element x satisfying P to 'smaller' y also satisfying P +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Induction.InfiniteDescent where + +open import Data.Nat.Base as ℕ using (ℕ; zero; suc) +open import Data.Nat.Properties as ℕ +open import Data.Product.Base using (_,_; proj₁; ∃-syntax; _×_) +open import Function.Base using (_∘_) +open import Induction.WellFounded + using (WellFounded; Acc; acc; acc-inverse; module Some) +open import Level using (Level) +open import Relation.Binary.Core using (Rel) +open import Relation.Binary.Construct.Closure.Transitive +open import Relation.Binary.PropositionalEquality.Core +open import Relation.Nullary.Negation.Core as Negation using (¬_) +open import Relation.Unary + using (Pred; ∁; _∩_; _⊆_; _⇒_; Universal; IUniversal; Stable; Empty) + +private + variable + a r ℓ : Level + A : Set a + f : ℕ → A + _<_ : Rel A r + P : Pred A ℓ + +------------------------------------------------------------------------ +-- Definitions + +InfiniteDescendingSequence : Rel A r → (ℕ → A) → Set _ +InfiniteDescendingSequence _<_ f = ∀ n → f (suc n) < f n + +InfiniteDescendingSequenceFrom : Rel A r → (ℕ → A) → Pred A _ +InfiniteDescendingSequenceFrom _<_ f x = f zero ≡ x × InfiniteDescendingSequence _<_ f + +InfiniteDescendingSequence⁺ : Rel A r → (ℕ → A) → Set _ +InfiniteDescendingSequence⁺ _<_ f = ∀ {m n} → m ℕ.< n → TransClosure _<_ (f n) (f m) + +InfiniteDescendingSequenceFrom⁺ : Rel A r → (ℕ → A) → Pred A _ +InfiniteDescendingSequenceFrom⁺ _<_ f x = f zero ≡ x × InfiniteDescendingSequence⁺ _<_ f + +DescentFrom : Rel A r → Pred A ℓ → Pred A _ +DescentFrom _<_ P x = P x → ∃[ y ] y < x × P y + +Descent : Rel A r → Pred A ℓ → Set _ +Descent _<_ P = ∀ {x} → DescentFrom _<_ P x + +InfiniteDescentFrom : Rel A r → Pred A ℓ → Pred A _ +InfiniteDescentFrom _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom _<_ f x × ∀ z → P (f z) + +InfiniteDescent : Rel A r → Pred A ℓ → Set _ +InfiniteDescent _<_ P = ∀ {x} → InfiniteDescentFrom _<_ P x + +InfiniteDescentFrom⁺ : Rel A r → Pred A ℓ → Pred A _ +InfiniteDescentFrom⁺ _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom⁺ _<_ f x × ∀ z → P (f z) + +InfiniteDescent⁺ : Rel A r → Pred A ℓ → Set _ +InfiniteDescent⁺ _<_ P = ∀ {x} → InfiniteDescentFrom⁺ _<_ P x + +NoSmallestCounterExample : Rel A r → Pred A ℓ → Set _ +NoSmallestCounterExample _<_ P = ∀ {x} → Acc _<_ x → DescentFrom (TransClosure _<_) (∁ P) x + +------------------------------------------------------------------------ +-- We can swap between transitively closed and non-transitively closed +-- definitions + +sequence⁺ : InfiniteDescendingSequence (TransClosure _<_) f → + InfiniteDescendingSequence⁺ _<_ f +sequence⁺ {_<_ = _<_} {f = f} seq[f] = seq⁺[f]′ ∘ ℕ.<⇒<′ + where + seq⁺[f]′ : ∀ {m n} → m ℕ.<′ n → TransClosure _<_ (f n) (f m) + seq⁺[f]′ ℕ.<′-base = seq[f] _ + seq⁺[f]′ (ℕ.<′-step m<′n) = seq[f] _ ++ seq⁺[f]′ m<′n + +sequence⁻ : InfiniteDescendingSequence⁺ _<_ f → + InfiniteDescendingSequence (TransClosure _<_) f +sequence⁻ seq[f] = seq[f] ∘ n<1+n + +------------------------------------------------------------------------ +-- Results about unrestricted descent + +descent+acc⇒infiniteDescentFrom : Descent _<_ P → (Acc _<_) ⊆ (InfiniteDescentFrom _<_ P) +descent+acc⇒infiniteDescentFrom {_<_ = _<_} {P = P} descent {x} = + Some.wfRec (InfiniteDescentFrom _<_ P) rec x + where + rec : _ + rec y rec[y] py + with z , z Date: Sun, 25 Feb 2024 21:02:02 +0800 Subject: [PATCH 55/57] Touch up names --- CHANGELOG.md | 6 +++--- src/Induction/InfiniteDescent.agda | 8 ++++---- 2 files changed, 7 insertions(+), 7 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 26fd5d190e..8ba32ad546 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -48,9 +48,9 @@ New modules * `Algebra.Module.Bundles.Raw`: raw bundles for module-like algebraic structures -* The 'no infinite descent' principle for (accessible elements of) well-founded relations: - ``` - Induction.NoInfiniteDescent +* Consequences of 'infinite descent' for (accessible elements of) well-founded relations: + ```agda + Induction.InfiniteDescent ``` * Nagata's construction of the "idealization of a module": diff --git a/src/Induction/InfiniteDescent.agda b/src/Induction/InfiniteDescent.agda index e0bd372102..2889200e37 100644 --- a/src/Induction/InfiniteDescent.agda +++ b/src/Induction/InfiniteDescent.agda @@ -178,10 +178,10 @@ module _ (accDescent⁺ : Acc _<_ ⊆ DescentFrom (TransClosure _<_) P) where module _ (stable : Stable P) where - acc⇒noSmallestCounterExample : NoSmallestCounterExample _<_ P → Acc _<_ ⊆ P - acc⇒noSmallestCounterExample noSmallest = + noSmallestCounterExample+acc⇒holds : NoSmallestCounterExample _<_ P → Acc _<_ ⊆ P + noSmallestCounterExample+acc⇒holds noSmallest = stable _ ∘ accDescent⁺+acc⇒notHold noSmallest - wf⇒noSmallestCounterExample : WellFounded _<_ → NoSmallestCounterExample _<_ P → Universal P - wf⇒noSmallestCounterExample wf noSmallest = + noSmallestCounterExample+wf⇒universal : NoSmallestCounterExample _<_ P → WellFounded _<_ → Universal P + noSmallestCounterExample+wf⇒universal noSmallest wf = stable _ ∘ accDescent⁺+wf⇒empty noSmallest wf From 88452ff4cd37349cd237741c26b7d6421119b4a5 Mon Sep 17 00:00:00 2001 From: MatthewDaggitt Date: Sun, 25 Feb 2024 21:03:31 +0800 Subject: [PATCH 56/57] Reintroduce anonymous module for descent --- src/Induction/InfiniteDescent.agda | 71 +++++++++++++++--------------- 1 file changed, 36 insertions(+), 35 deletions(-) diff --git a/src/Induction/InfiniteDescent.agda b/src/Induction/InfiniteDescent.agda index 2889200e37..4893e36ec4 100644 --- a/src/Induction/InfiniteDescent.agda +++ b/src/Induction/InfiniteDescent.agda @@ -87,43 +87,44 @@ sequence⁻ seq[f] = seq[f] ∘ n<1+n ------------------------------------------------------------------------ -- Results about unrestricted descent -descent+acc⇒infiniteDescentFrom : Descent _<_ P → (Acc _<_) ⊆ (InfiniteDescentFrom _<_ P) -descent+acc⇒infiniteDescentFrom {_<_ = _<_} {P = P} descent {x} = - Some.wfRec (InfiniteDescentFrom _<_ P) rec x - where - rec : _ - rec y rec[y] py - with z , z Date: Mon, 26 Feb 2024 13:32:26 +0000 Subject: [PATCH 57/57] cosmetics asper final comment --- src/Induction/InfiniteDescent.agda | 69 +++++++++++++++--------------- 1 file changed, 35 insertions(+), 34 deletions(-) diff --git a/src/Induction/InfiniteDescent.agda b/src/Induction/InfiniteDescent.agda index 4893e36ec4..ed25c880d5 100644 --- a/src/Induction/InfiniteDescent.agda +++ b/src/Induction/InfiniteDescent.agda @@ -54,13 +54,13 @@ Descent : Rel A r → Pred A ℓ → Set _ Descent _<_ P = ∀ {x} → DescentFrom _<_ P x InfiniteDescentFrom : Rel A r → Pred A ℓ → Pred A _ -InfiniteDescentFrom _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom _<_ f x × ∀ z → P (f z) +InfiniteDescentFrom _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom _<_ f x × ∀ n → P (f n) InfiniteDescent : Rel A r → Pred A ℓ → Set _ InfiniteDescent _<_ P = ∀ {x} → InfiniteDescentFrom _<_ P x InfiniteDescentFrom⁺ : Rel A r → Pred A ℓ → Pred A _ -InfiniteDescentFrom⁺ _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom⁺ _<_ f x × ∀ z → P (f z) +InfiniteDescentFrom⁺ _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom⁺ _<_ f x × ∀ n → P (f n) InfiniteDescent⁺ : Rel A r → Pred A ℓ → Set _ InfiniteDescent⁺ _<_ P = ∀ {x} → InfiniteDescentFrom⁺ _<_ P x @@ -89,8 +89,8 @@ sequence⁻ seq[f] = seq[f] ∘ n<1+n module _ (descent : Descent _<_ P) where - descent+acc⇒infiniteDescentFrom : (Acc _<_) ⊆ (InfiniteDescentFrom _<_ P) - descent+acc⇒infiniteDescentFrom {x} = + descent∧acc⇒infiniteDescentFrom : (Acc _<_) ⊆ (InfiniteDescentFrom _<_ P) + descent∧acc⇒infiniteDescentFrom {x} = Some.wfRec (InfiniteDescentFrom _<_ P) rec x where rec : _ @@ -114,17 +114,17 @@ module _ (descent : Descent _<_ P) where Π[P∘h] zero rewrite g0≡z = py Π[P∘h] (suc n) = Π[P∘g] n - descent+wf⇒infiniteDescent : WellFounded _<_ → InfiniteDescent _<_ P - descent+wf⇒infiniteDescent wf = descent+acc⇒infiniteDescentFrom (wf _) + descent∧wf⇒infiniteDescent : WellFounded _<_ → InfiniteDescent _<_ P + descent∧wf⇒infiniteDescent wf = descent∧acc⇒infiniteDescentFrom (wf _) - descent+acc⇒notHold : Acc _<_ ⊆ ∁ P - descent+acc⇒notHold {x} = Some.wfRec (∁ P) rec x + descent∧acc⇒unsatisfiable : Acc _<_ ⊆ ∁ P + descent∧acc⇒unsatisfiable {x} = Some.wfRec (∁ P) rec x where rec : _ rec y rec[y] py = let z , z