From 43c73161f63d20b891a3ea74127eb86ed232f232 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Tue, 8 Jul 2025 08:29:15 +0100 Subject: [PATCH 1/2] add: rule of signs for `RingWithoutOne` --- CHANGELOG.md | 10 +++++++--- src/Algebra/Properties/RingWithoutOne.agda | 20 +++++++++++++++----- 2 files changed, 22 insertions(+), 8 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 4b10f1a54e..1555c53894 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -254,9 +254,13 @@ Additions to existing modules ∣ˡ-preorder : Preorder a ℓ _ ``` -* In `Algebra.Properties.Semigroup` adding consequences for associativity for semigroups +* In `Algebra.Properties.RingWithoutOne`: + ```agda + [-x][-y]≈xy : ∀ x y → - x * - y ≈ x * y + ``` -``` +* In `Algebra.Properties.Semigroup`, consequences for associativity for semigroups: + ``` uv≈w⇒xu∙v≈xw : ∀ x → (x ∙ u) ∙ v ≈ x ∙ w uv≈w⇒u∙vx≈wx : ∀ x → u ∙ (v ∙ x) ≈ w ∙ x uv≈w⇒u[vx∙y]≈w∙xy : ∀ x y → u ∙ ((v ∙ x) ∙ y) ≈ w ∙ (x ∙ y) @@ -275,7 +279,7 @@ Additions to existing modules uv≈wx⇒yu∙v≈yw∙x : ∀ y → (y ∙ u) ∙ v ≈ (y ∙ w) ∙ x uv≈wx⇒u∙vy≈w∙xy : ∀ y → u ∙ (v ∙ y) ≈ w ∙ (x ∙ y) uv≈wx⇒yu∙vz≈yw∙xz : ∀ y z → (y ∙ u) ∙ (v ∙ z) ≈ (y ∙ w) ∙ (x ∙ z) -``` + ``` * In `Algebra.Properties.Semigroup.Divisibility`: ```agda diff --git a/src/Algebra/Properties/RingWithoutOne.agda b/src/Algebra/Properties/RingWithoutOne.agda index 5dda568cbc..d45edeaf0b 100644 --- a/src/Algebra/Properties/RingWithoutOne.agda +++ b/src/Algebra/Properties/RingWithoutOne.agda @@ -17,7 +17,7 @@ open import Function.Base using (_$_) open import Relation.Binary.Reasoning.Setoid setoid ------------------------------------------------------------------------ --- Export properties of abelian groups +-- Re-export abelian group properties for addition open AbelianGroupProperties +-abelianGroup public renaming @@ -36,6 +36,12 @@ open AbelianGroupProperties +-abelianGroup public ; ⁻¹-∙-comm to -‿+-comm ) +x+x≈x⇒x≈0 : ∀ x → x + x ≈ x → x ≈ 0# +x+x≈x⇒x≈0 x eq = +-identityˡ-unique x x eq + +------------------------------------------------------------------------ +-- Consequences of distributivity + -‿distribˡ-* : ∀ x y → - (x * y) ≈ - x * y -‿distribˡ-* x y = sym $ begin - x * y ≈⟨ +-identityʳ (- x * y) ⟨ @@ -58,17 +64,21 @@ open AbelianGroupProperties +-abelianGroup public - (x * y) + 0# ≈⟨ +-identityʳ (- (x * y)) ⟩ - (x * y) ∎ -x+x≈x⇒x≈0 : ∀ x → x + x ≈ x → x ≈ 0# -x+x≈x⇒x≈0 x eq = +-identityˡ-unique x x eq +[-x][-y]≈xy : ∀ x y → - x * - y ≈ x * y +[-x][-y]≈xy x y = begin + - x * - y ≈⟨ -‿distribˡ-* x (- y) ⟨ + - (x * - y) ≈⟨ -‿cong (-‿distribʳ-* x y) ⟨ + - (- (x * y)) ≈⟨ -‿involutive (x * y) ⟩ + x * y ∎ x[y-z]≈xy-xz : ∀ x y z → x * (y - z) ≈ x * y - x * z x[y-z]≈xy-xz x y z = begin x * (y - z) ≈⟨ distribˡ x y (- z) ⟩ - x * y + x * - z ≈⟨ +-congˡ (sym (-‿distribʳ-* x z)) ⟩ + x * y + x * - z ≈⟨ +-congˡ (-‿distribʳ-* x z) ⟨ x * y - x * z ∎ [y-z]x≈yx-zx : ∀ x y z → (y - z) * x ≈ (y * x) - (z * x) [y-z]x≈yx-zx x y z = begin (y - z) * x ≈⟨ distribʳ x y (- z) ⟩ - y * x + - z * x ≈⟨ +-congˡ (sym (-‿distribˡ-* z x)) ⟩ + y * x + - z * x ≈⟨ +-congˡ (-‿distribˡ-* z x) ⟨ y * x - z * x ∎ From 6b84352ddee22623c25fcfa4c217f2f54874050b Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sun, 20 Jul 2025 20:38:39 +0100 Subject: [PATCH 2/2] reset: `CHANGELOG` for v2.4 --- CHANGELOG.md | 528 +-------------------------------------------------- 1 file changed, 2 insertions(+), 526 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 1555c53894..0eaaa0c00e 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1,7 +1,7 @@ -Version 2.3-dev +Version 2.4-dev =============== -The library has been tested using Agda 2.7.0 and 2.7.0.1. +The library has been tested using Agda 2.8.0. Highlights ---------- @@ -9,550 +9,26 @@ Highlights Bug-fixes --------- -* In `Algebra.Apartness.Structures`, renamed `sym` from `IsApartnessRelation` - to `#-sym` in order to avoid overloaded projection. - `irrefl` and `cotrans` are similarly renamed for the sake of consistency. - -* In `Algebra.Definitions.RawMagma` and `Relation.Binary.Construct.Interior.Symmetric`, - the record constructors `_,_` incorrectly had no declared fixity. They have been given - the fixity `infixr 4 _,_`, consistent with that of `Data.Product.Base`. - -* In `Data.Product.Function.Dependent.Setoid`, `left-inverse` defined a - `RightInverse`. - This has been deprecated in favor or `rightInverse`, and a corrected (and - correctly-named) function `leftInverse` has been added. - -* The implementation of `_IsRelatedTo_` in `Relation.Binary.Reasoning.Base.Partial` - has been modified to correctly support equational reasoning at the beginning and the end. - The detail of this issue is described in [#2677](https://github.com/agda/agda-stdlib/pull/2677). Since the names of constructors - of `_IsRelatedTo_` are changed and the reduction behaviour of reasoning steps - are changed, this modification is non-backwards compatible. - Non-backwards compatible changes -------------------------------- -* The implementation of `≤-total` in `Data.Nat.Properties` has been altered - to use operations backed by primitives, rather than recursion, making it - significantly faster. However, its reduction behaviour on open terms may have - changed. - Minor improvements ------------------ -* Moved the concept `Irrelevant` of irrelevance (h-proposition) from `Relation.Nullary` - to its own dedicated module `Relation.Nullary.Irrelevant`. - Deprecated modules ------------------ Deprecated names ---------------- -* In `Algebra.Definitions.RawMagma`: - ```agda - _∣∣_ ↦ _∥_ - _∤∤_ ↦ _∦_ - ``` - -* In `Algebra.Lattice.Properties.BooleanAlgebra` - ```agda - ⊥≉⊤ ↦ ¬⊥≈⊤ - ⊤≉⊥ ↦ ¬⊤≈⊥ - ``` - -* In `Algebra.Module.Consequences` - ```agda - *ₗ-assoc+comm⇒*ᵣ-assoc ↦ *ₗ-assoc∧comm⇒*ᵣ-assoc - *ₗ-assoc+comm⇒*ₗ-*ᵣ-assoc ↦ *ₗ-assoc∧comm⇒*ₗ-*ᵣ-assoc - *ᵣ-assoc+comm⇒*ₗ-assoc ↦ *ᵣ-assoc∧comm⇒*ₗ-assoc - *ₗ-assoc+comm⇒*ₗ-*ᵣ-assoc ↦ *ₗ-assoc∧comm⇒*ₗ-*ᵣ-assoc - ``` - -* In `Algebra.Modules.Structures.IsLeftModule`: - ```agda - uniqueˡ‿⁻ᴹ ↦ Algebra.Module.Properties.LeftModule.inverseˡ-uniqueᴹ - uniqueʳ‿⁻ᴹ ↦ Algebra.Module.Properties.LeftModule.inverseʳ-uniqueᴹ - ``` - -* In `Algebra.Modules.Structures.IsRightModule`: - ```agda - uniqueˡ‿⁻ᴹ ↦ Algebra.Module.Properties.RightModule.inverseˡ-uniqueᴹ - uniqueʳ‿⁻ᴹ ↦ Algebra.Module.Properties.RightModule.inverseʳ-uniqueᴹ - ``` - -* In `Algebra.Properties.Magma.Divisibility`: - ```agda - ∣∣-sym ↦ ∥-sym - ∣∣-respˡ-≈ ↦ ∥-respˡ-≈ - ∣∣-respʳ-≈ ↦ ∥-respʳ-≈ - ∣∣-resp-≈ ↦ ∥-resp-≈ - ∤∤-sym -≈ ↦ ∦-sym - ∤∤-respˡ-≈ ↦ ∦-respˡ-≈ - ∤∤-respʳ-≈ ↦ ∦-respʳ-≈ - ∤∤-resp-≈ ↦ ∦-resp-≈ - ∣-respʳ-≈ ↦ ∣ʳ-respʳ-≈ - ∣-respˡ-≈ ↦ ∣ʳ-respˡ-≈ - ∣-resp-≈ ↦ ∣ʳ-resp-≈ - x∣yx ↦ x∣ʳyx - xy≈z⇒y∣z ↦ xy≈z⇒y∣ʳz - ``` - -* In `Algebra.Properties.Monoid.Divisibility`: - ```agda - ∣∣-refl ↦ ∥-refl - ∣∣-reflexive ↦ ∥-reflexive - ∣∣-isEquivalence ↦ ∥-isEquivalence - ε∣_ ↦ ε∣ʳ_ - ∣-refl ↦ ∣ʳ-refl - ∣-reflexive ↦ ∣ʳ-reflexive - ∣-isPreorder ↦ ∣ʳ-isPreorder - ∣-preorder ↦ ∣ʳ-preorder - ``` - -* In `Algebra.Properties.Semigroup.Divisibility`: - ```agda - ∣∣-trans ↦ ∥-trans - ∣-trans ↦ ∣ʳ-trans - ``` - -* In `Algebra.Structures.Group`: - ```agda - uniqueˡ-⁻¹ ↦ Algebra.Properties.Group.inverseˡ-unique - uniqueʳ-⁻¹ ↦ Algebra.Properties.Group.inverseʳ-unique - ``` - -* In `Data.List.Base`: - ```agda - and ↦ Data.Bool.ListAction.and - or ↦ Data.Bool.ListAction.or - any ↦ Data.Bool.ListAction.any - all ↦ Data.Bool.ListAction.all - sum ↦ Data.Nat.ListAction.sum - product ↦ Data.Nat.ListAction.product - ``` - -* In `Data.List.Properties`: - ```agda - sum-++ ↦ Data.Nat.ListAction.Properties.sum-++ - ∈⇒∣product ↦ Data.Nat.ListAction.Properties.∈⇒∣product - product≢0 ↦ Data.Nat.ListAction.Properties.product≢0 - ∈⇒≤product ↦ Data.Nat.ListAction.Properties.∈⇒≤product - ∷-ʳ++-eqFree ↦ Data.List.Properties.ʳ++-ʳ++-eqFree - ``` - -* In `Data.List.Relation.Binary.Permutation.Propositional.Properties`: - ```agda - sum-↭ ↦ Data.Nat.ListAction.Properties.sum-↭ - product-↭ ↦ Data.Nat.ListAction.Properties.product-↭ - ``` - -* In `Data.Product.Function.Dependent.Setoid`: - ```agda - left-inverse ↦ rightInverse - ``` - -* In `Data.Product.Nary.NonDependent`: - ```agda - Allₙ ↦ Pointwiseₙ - ``` - New modules ----------- -* `Algebra.Module.Properties.{Bimodule|LeftModule|RightModule}`. - -* `Algebra.Morphism.Construct.DirectProduct`. - -* `Data.List.Base.{and|or|any|all}` have been lifted out into `Data.Bool.ListAction`. - -* `Data.List.Base.{sum|product}` and their properties have been lifted out into `Data.Nat.ListAction` and `Data.Nat.ListAction.Properties`. - -* `Data.List.Relation.Binary.Prefix.Propositional.Properties` showing the equivalence to left divisibility induced by the list monoid. - -* `Data.List.Relation.Binary.Suffix.Propositional.Properties` showing the equivalence to right divisibility induced by the list monoid. - -* `Data.List.Sort.InsertionSort.{agda|Base|Properties}` defines insertion sort and proves properties of insertion sort such as Sorted and Permutation properties. - -* `Data.List.Sort.MergenSort.{agda|Base|Properties}` is a refactor of the previous `Data.List.Sort.MergenSort`. - -* `Data.Sign.Show` to show a sign. - -* `Relation.Binary.Morphism.Construct.Product` to plumb in the (categorical) product structure on `RawSetoid`. - -* `Relation.Binary.Properties.PartialSetoid` to systematise properties of a PER - -* `Relation.Nullary.Recomputable.Core` - Additions to existing modules ----------------------------- -* In `Algebra.Consequences.Base`: - ```agda - module Congruence (_≈_ : Rel A ℓ) (cong : Congruent₂ _≈_ _∙_) (refl : Reflexive _≈_) - where - ∙-congˡ : LeftCongruent _≈_ _∙_ - ∙-congʳ : RightCongruent _≈_ _∙_ - ``` - -* In `Algebra.Consequences.Setoid`: - ```agda - module Congruence (cong : Congruent₂ _≈_ _∙_) where - ∙-congˡ : LeftCongruent _≈_ _∙_ - ∙-congʳ : RightCongruent _≈_ _∙_ - ``` - -* In `Algebra.Construct.Pointwise`: - ```agda - isNearSemiring : IsNearSemiring _≈_ _+_ _*_ 0# → - IsNearSemiring (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₀ 0#) - isSemiringWithoutOne : IsSemiringWithoutOne _≈_ _+_ _*_ 0# → - IsSemiringWithoutOne (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₀ 0#) - isCommutativeSemiringWithoutOne : IsCommutativeSemiringWithoutOne _≈_ _+_ _*_ 0# → - IsCommutativeSemiringWithoutOne (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₀ 0#) - isCommutativeSemiring : IsCommutativeSemiring _≈_ _+_ _*_ 0# 1# → - IsCommutativeSemiring (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₀ 0#) (lift₀ 1#) - isIdempotentSemiring : IsIdempotentSemiring _≈_ _+_ _*_ 0# 1# → - IsIdempotentSemiring (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₀ 0#) (lift₀ 1#) - isKleeneAlgebra : IsKleeneAlgebra _≈_ _+_ _*_ _⋆ 0# 1# → - IsKleeneAlgebra (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₁ _⋆) (lift₀ 0#) (lift₀ 1#) - isQuasiring : IsQuasiring _≈_ _+_ _*_ 0# 1# → - IsQuasiring (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₀ 0#) (lift₀ 1#) - isCommutativeRing : IsCommutativeRing _≈_ _+_ _*_ -_ 0# 1# → - IsCommutativeRing (liftRel _≈_) (lift₂ _+_) (lift₂ _*_) (lift₁ -_) (lift₀ 0#) (lift₀ 1#) - commutativeMonoid : CommutativeMonoid c ℓ → CommutativeMonoid (a ⊔ c) (a ⊔ ℓ) - nearSemiring : NearSemiring c ℓ → NearSemiring (a ⊔ c) (a ⊔ ℓ) - semiringWithoutOne : SemiringWithoutOne c ℓ → SemiringWithoutOne (a ⊔ c) (a ⊔ ℓ) - commutativeSemiringWithoutOne : CommutativeSemiringWithoutOne c ℓ → CommutativeSemiringWithoutOne (a ⊔ c) (a ⊔ ℓ) - commutativeSemiring : CommutativeSemiring c ℓ → CommutativeSemiring (a ⊔ c) (a ⊔ ℓ) - idempotentSemiring : IdempotentSemiring c ℓ → IdempotentSemiring (a ⊔ c) (a ⊔ ℓ) - kleeneAlgebra : KleeneAlgebra c ℓ → KleeneAlgebra (a ⊔ c) (a ⊔ ℓ) - quasiring : Quasiring c ℓ → Quasiring (a ⊔ c) (a ⊔ ℓ) - commutativeRing : CommutativeRing c ℓ → CommutativeRing (a ⊔ c) (a ⊔ ℓ) - ``` - -* In `Algebra.Modules.Properties`: - ```agda - inverseˡ-uniqueᴹ : x +ᴹ y ≈ 0ᴹ → x ≈ -ᴹ y - inverseʳ-uniqueᴹ : x +ᴹ y ≈ 0ᴹ → y ≈ -ᴹ x - ``` - -* In `Algebra.Properties.Magma.Divisibility`: - ```agda - ∣ˡ-respʳ-≈ : _∣ˡ_ Respectsʳ _≈_ - ∣ˡ-respˡ-≈ : _∣ˡ_ Respectsˡ _≈_ - ∣ˡ-resp-≈ : _∣ˡ_ Respects₂ _≈_ - x∣ˡxy : ∀ x y → x ∣ˡ x ∙ y - xy≈z⇒x∣ˡz : ∀ x y {z} → x ∙ y ≈ z → x ∣ˡ z - ``` - -* In `Algebra.Properties.Monoid.Divisibility`: - ```agda - ε∣ˡ_ : ∀ x → ε ∣ˡ x - ∣ˡ-refl : Reflexive _∣ˡ_ - ∣ˡ-reflexive : _≈_ ⇒ _∣ˡ_ - ∣ˡ-isPreorder : IsPreorder _≈_ _∣ˡ_ - ∣ˡ-preorder : Preorder a ℓ _ - ``` - * In `Algebra.Properties.RingWithoutOne`: ```agda [-x][-y]≈xy : ∀ x y → - x * - y ≈ x * y ``` -* In `Algebra.Properties.Semigroup`, consequences for associativity for semigroups: - ``` - uv≈w⇒xu∙v≈xw : ∀ x → (x ∙ u) ∙ v ≈ x ∙ w - uv≈w⇒u∙vx≈wx : ∀ x → u ∙ (v ∙ x) ≈ w ∙ x - uv≈w⇒u[vx∙y]≈w∙xy : ∀ x y → u ∙ ((v ∙ x) ∙ y) ≈ w ∙ (x ∙ y) - uv≈w⇒x[uv∙y]≈x∙wy : ∀ x y → x ∙ (u ∙ (v ∙ y)) ≈ x ∙ (w ∙ y) - uv≈w⇒[x∙yu]v≈x∙yw : ∀ x y → (x ∙ (y ∙ u)) ∙ v ≈ x ∙ (y ∙ w) - uv≈w⇒[xu∙v]y≈x∙wy : ∀ x y → ((x ∙ u) ∙ v) ∙ y ≈ x ∙ (w ∙ y) - uv≈w⇒[xy∙u]v≈x∙yw : ∀ x y → ((x ∙ y) ∙ u) ∙ v ≈ x ∙ (y ∙ w) - uv≈w⇒xu∙vy≈x∙wy : ∀ x y → (x ∙ u) ∙ (v ∙ y) ≈ x ∙ (w ∙ y) - uv≈w⇒xy≈z⇒u[vx∙y]≈wz : ∀ z → x ∙ y ≈ z → u ∙ ((v ∙ x) ∙ y) ≈ w ∙ z - uv≈w⇒x∙wy≈x∙[u∙vy] : x ∙ (w ∙ y) ≈ x ∙ (u ∙ (v ∙ y)) - [uv∙w]x≈u[vw∙x] : ((u ∙ v) ∙ w) ∙ x ≈ u ∙ ((v ∙ w) ∙ x) - [uv∙w]x≈u[v∙wx] : ((u ∙ v) ∙ w) ∙ x ≈ u ∙ (v ∙ (w ∙ x)) - [u∙vw]x≈uv∙wx : (u ∙ (v ∙ w)) ∙ x ≈ (u ∙ v) ∙ (w ∙ x) - [u∙vw]x≈u[v∙wx] : (u ∙ (v ∙ w)) ∙ x ≈ u ∙ (v ∙ (w ∙ x)) - uv∙wx≈u[vw∙x] : (u ∙ v) ∙ (w ∙ x) ≈ u ∙ ((v ∙ w) ∙ x) - uv≈wx⇒yu∙v≈yw∙x : ∀ y → (y ∙ u) ∙ v ≈ (y ∙ w) ∙ x - uv≈wx⇒u∙vy≈w∙xy : ∀ y → u ∙ (v ∙ y) ≈ w ∙ (x ∙ y) - uv≈wx⇒yu∙vz≈yw∙xz : ∀ y z → (y ∙ u) ∙ (v ∙ z) ≈ (y ∙ w) ∙ (x ∙ z) - ``` - -* In `Algebra.Properties.Semigroup.Divisibility`: - ```agda - ∣ˡ-trans : Transitive _∣ˡ_ - x∣ʳy⇒x∣ʳzy : x ∣ʳ y → x ∣ʳ z ∙ y - x∣ʳy⇒xz∣ʳyz : x ∣ʳ y → x ∙ z ∣ʳ y ∙ z - x∣ˡy⇒zx∣ˡzy : x ∣ˡ y → z ∙ x ∣ˡ z ∙ y - x∣ˡy⇒x∣ˡyz : x ∣ˡ y → x ∣ˡ y ∙ z - ``` - -* In `Algebra.Properties.CommutativeSemigroup.Divisibility`: - ```agda - ∙-cong-∣ : x ∣ y → a ∣ b → x ∙ a ∣ y ∙ b - ``` - -* In `Data.Bool.Properties`: - ```agda - if-eta : ∀ b → (if b then x else x) ≡ x - if-idem-then : ∀ b → (if b then (if b then x else y) else y) ≡ (if b then x else y) - if-idem-else : ∀ b → (if b then x else (if b then x else y)) ≡ (if b then x else y) - if-swap-then : ∀ b c → (if b then (if c then x else y) else y) ≡ (if c then (if b then x else y) else y) - if-swap-else : ∀ b c → (if b then x else (if c then x else y)) ≡ (if c then x else (if b then x else y)) - if-not : ∀ b → (if not b then x else y) ≡ (if b then y else x) - if-∧ : ∀ b → (if b ∧ c then x else y) ≡ (if b then (if c then x else y) else y) - if-∨ : ∀ b → (if b ∨ c then x else y) ≡ (if b then x else (if c then x else y)) - if-xor : ∀ b → (if b xor c then x else y) ≡ (if b then (if c then y else x) else (if c then x else y)) - if-cong : b ≡ c → (if b then x else y) ≡ (if c then x else y) - if-cong-then : ∀ b → x ≡ z → (if b then x else y) ≡ (if b then z else y) - if-cong-else : ∀ b → y ≡ z → (if b then x else y) ≡ (if b then x else z) - if-cong₂ : ∀ b → x ≡ z → y ≡ w → (if b then x else y) ≡ (if b then z else w) - ``` - -* In `Data.Fin.Base`: - ```agda - _≰_ : Rel (Fin n) 0ℓ - _≮_ : Rel (Fin n) 0ℓ - lower : ∀ (i : Fin m) → .(toℕ i ℕ.< n) → Fin n - ``` - -* In `Data.Fin.Permutation`: - ```agda - cast-id : .(m ≡ n) → Permutation m n - swap : Permutation m n → Permutation (2+ m) (2+ n) - ``` - -* In `Data.Fin.Properties`: - ```agda - cast-involutive : .(eq₁ : m ≡ n) .(eq₂ : n ≡ m) → ∀ k → cast eq₁ (cast eq₂ k) ≡ k - inject!-injective : Injective _≡_ _≡_ inject! - inject!-< : (k : Fin′ i) → inject! k < i - lower-injective : lower i i